src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML
author blanchet
Fri, 27 Aug 2010 18:04:49 +0200
changeset 38889 d0e3f68dde63
parent 38829 c18e8f90f4dc
child 38891 6e47e54214b8
permissions -rw-r--r--
fiddle with the relevance filter
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML
38027
505657ddb047 standardize "Author" tags
blanchet
parents: 37995
diff changeset
     2
    Author:     Jia Meng, Cambridge University Computer Laboratory and NICTA
36393
be73a2b2443b support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents: 36227
diff changeset
     3
    Author:     Jasmin Blanchette, TU Muenchen
33309
5f67433e6dd8 proper header;
wenzelm
parents: 33306
diff changeset
     4
*)
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
     5
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
     6
signature SLEDGEHAMMER_FACT_FILTER =
16802
6eeee59dac4c use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents: 16767
diff changeset
     7
sig
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
     8
  datatype locality = General | Theory | Local | Chained
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
     9
35966
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    10
  type relevance_override =
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    11
    {add: Facts.ref list,
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    12
     del: Facts.ref list,
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    13
     only: bool}
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    14
37616
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
    15
  val trace : bool Unsynchronized.ref
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
    16
  val name_thm_pairs_from_ref :
38696
4c6b65d6a135 quote facts whose names collide with a keyword or command name (cf. "subclass" in "Jinja/J/TypeSafe.thy")
blanchet
parents: 38692
diff changeset
    17
    Proof.context -> unit Symtab.table -> thm list -> Facts.ref
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
    18
    -> ((string * locality) * thm) list
37347
635425a442e8 show more respect for user-specified facts, even if they could lead to unsound proofs + don't throw away "unsound" theorems in "full_type" mode, since they are then sound
blanchet
parents: 37345
diff changeset
    19
  val relevant_facts :
38745
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
    20
    bool -> real * real -> int -> bool -> relevance_override
37995
06f02b15ef8a generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents: 37626
diff changeset
    21
    -> Proof.context * (thm list * 'a) -> term list -> term
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
    22
    -> ((string * locality) * thm) list
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    23
end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    24
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
    25
structure Sledgehammer_Fact_Filter : SLEDGEHAMMER_FACT_FILTER =
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    26
struct
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    27
38652
e063be321438 perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents: 38644
diff changeset
    28
open Sledgehammer_Util
e063be321438 perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents: 38644
diff changeset
    29
37616
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
    30
val trace = Unsynchronized.ref false
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
    31
fun trace_msg msg = if !trace then tracing (msg ()) else ()
35826
1590abc3d42a renamed Sledgehammer structures
blanchet
parents: 35825
diff changeset
    32
38828
91ad85f962c4 turn off experimental feature per default + avoid exception on "theory constant"
blanchet
parents: 38827
diff changeset
    33
(* experimental feature *)
38889
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
    34
val term_patterns = false
38827
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
    35
37580
c2c1caff5dea got rid of "respect_no_atp" option, which even I don't use
blanchet
parents: 37578
diff changeset
    36
val respect_no_atp = true
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
    37
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
    38
datatype locality = General | Theory | Local | Chained
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
    39
35966
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    40
type relevance_override =
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    41
  {add: Facts.ref list,
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    42
   del: Facts.ref list,
f8c738abaed8 honor the newly introduced Sledgehammer parameters and fixed the parsing;
blanchet
parents: 35963
diff changeset
    43
   only: bool}
21070
0a898140fea2 Added more debugging info
paulson
parents: 20995
diff changeset
    44
37616
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
    45
val sledgehammer_prefix = "Sledgehammer" ^ Long_Name.separator
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
    46
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
    47
fun repair_name reserved multi j name =
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
    48
  (name |> Symtab.defined reserved name ? quote) ^
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
    49
  (if multi then "(" ^ Int.toString j ^ ")" else "")
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
    50
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
    51
fun name_thm_pairs_from_ref ctxt reserved chained_ths xref =
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
    52
  let
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
    53
    val ths = ProofContext.get_fact ctxt xref
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
    54
    val name = Facts.string_of_ref xref
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
    55
    val multi = length ths > 1
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
    56
  in
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
    57
    (ths, (1, []))
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
    58
    |-> fold (fn th => fn (j, rest) =>
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
    59
                 (j + 1, ((repair_name reserved multi j name,
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
    60
                          if member Thm.eq_thm chained_ths th then Chained
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
    61
                          else General), th) :: rest))
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
    62
    |> snd
38699
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
    63
  end
37616
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
    64
28477
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    65
(***************************************************************)
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    66
(* Relevance Filtering                                         *)
9339d4dcec8b version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents: 28065
diff changeset
    67
(***************************************************************)
19194
7681c04d8bff Merged res_atp_setup.ML into res_atp.ML.
mengj
parents: 18986
diff changeset
    68
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    69
(*** constants with types ***)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    70
38823
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
    71
(* An abstraction of Isabelle types and first-order terms *)
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
    72
datatype pattern = PVar | PApp of string * pattern list
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
    73
38823
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
    74
fun string_for_pattern PVar = "_"
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
    75
  | string_for_pattern (PApp (s, ps)) =
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
    76
    if null ps then s else s ^ string_for_patterns ps
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
    77
and string_for_patterns ps = "(" ^ commas (map string_for_pattern ps) ^ ")"
24287
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?*)
38824
f74513bbe627 cosmetics
blanchet
parents: 38823
diff changeset
    80
fun match_pattern (PVar, _) = true
f74513bbe627 cosmetics
blanchet
parents: 38823
diff changeset
    81
  | match_pattern (PApp _, PVar) = false
f74513bbe627 cosmetics
blanchet
parents: 38823
diff changeset
    82
  | match_pattern (PApp (s, ps), PApp (t, qs)) =
f74513bbe627 cosmetics
blanchet
parents: 38823
diff changeset
    83
    s = t andalso match_patterns (ps, qs)
f74513bbe627 cosmetics
blanchet
parents: 38823
diff changeset
    84
and match_patterns (_, []) = true
f74513bbe627 cosmetics
blanchet
parents: 38823
diff changeset
    85
  | match_patterns ([], _) = false
f74513bbe627 cosmetics
blanchet
parents: 38823
diff changeset
    86
  | match_patterns (p :: ps, q :: qs) =
f74513bbe627 cosmetics
blanchet
parents: 38823
diff changeset
    87
    match_pattern (p, q) andalso match_patterns (ps, qs)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
    88
38823
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
    89
(* Is there a unifiable constant? *)
38827
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
    90
fun pconst_mem f consts (s, ps) =
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
    91
  exists (curry (match_patterns o f) ps)
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
    92
         (map snd (filter (curry (op =) s o fst) consts))
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
    93
fun pconst_hyper_mem f const_tab (s, ps) =
38823
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
    94
  exists (curry (match_patterns o f) ps) (these (Symtab.lookup const_tab s))
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
    95
38827
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
    96
fun ptype (Type (s, Ts)) = PApp (s, map ptype Ts)
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
    97
  | ptype (TFree (s, _)) = PApp (s, [])
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
    98
  | ptype (TVar _) = PVar
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
    99
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   100
fun pterm thy t =
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   101
  case strip_comb t of
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   102
    (Const x, ts) => PApp (pconst thy true x ts)
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   103
  | (Free x, ts) => PApp (pconst thy false x ts)
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   104
  | (Var x, []) => PVar
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   105
  | _ => PApp ("?", [])  (* equivalence class of higher-order constructs *)
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   106
(* Pairs a constant with the list of its type instantiations. *)
38827
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   107
and pconst_args thy const (s, T) ts =
38828
91ad85f962c4 turn off experimental feature per default + avoid exception on "theory constant"
blanchet
parents: 38827
diff changeset
   108
  (if const then map ptype (these (try (Sign.const_typargs thy) (s, T)))
91ad85f962c4 turn off experimental feature per default + avoid exception on "theory constant"
blanchet
parents: 38827
diff changeset
   109
   else []) @
38889
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   110
  (if term_patterns then map (pterm thy) ts else [])
38827
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   111
and pconst thy const (s, T) ts = (s, pconst_args thy const (s, T) ts)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   112
38827
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   113
fun string_for_hyper_pconst (s, pss) =
38823
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
   114
  s ^ "{" ^ commas (map string_for_patterns pss) ^ "}"
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   115
38816
21a6f261595e lower penalty for Skolem constants
blanchet
parents: 38752
diff changeset
   116
val abs_name = "Sledgehammer.abs"
38749
0d2f7f0614d1 add a penalty for lambda-abstractions;
blanchet
parents: 38747
diff changeset
   117
val skolem_prefix = "Sledgehammer.sko"
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   118
38822
aa0101e618e2 fix threshold computation + remove "op =" from relevant constants
blanchet
parents: 38821
diff changeset
   119
(* These are typically simplified away by "Meson.presimplify". Equality is
aa0101e618e2 fix threshold computation + remove "op =" from relevant constants
blanchet
parents: 38821
diff changeset
   120
   handled specially via "fequal". *)
38682
3a203da3f89b weed out junk in relevance filter
blanchet
parents: 38681
diff changeset
   121
val boring_consts =
38822
aa0101e618e2 fix threshold computation + remove "op =" from relevant constants
blanchet
parents: 38821
diff changeset
   122
  [@{const_name False}, @{const_name True}, @{const_name If}, @{const_name Let},
aa0101e618e2 fix threshold computation + remove "op =" from relevant constants
blanchet
parents: 38821
diff changeset
   123
   @{const_name "op ="}]
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   124
38823
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
   125
(* Add a pconstant to the table, but a [] entry means a standard
38819
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   126
   connective, which we ignore.*)
38824
f74513bbe627 cosmetics
blanchet
parents: 38823
diff changeset
   127
fun add_pconst_to_table also_skolem (c, ps) =
38819
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   128
  if member (op =) boring_consts c orelse
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   129
     (not also_skolem andalso String.isPrefix skolem_prefix c) then
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   130
    I
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   131
  else
38824
f74513bbe627 cosmetics
blanchet
parents: 38823
diff changeset
   132
    Symtab.map_default (c, [ps]) (insert (op =) ps)
38819
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   133
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   134
fun is_formula_type T = (T = HOLogic.boolT orelse T = propT)
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   135
38823
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
   136
fun get_pconsts thy also_skolems pos ts =
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   137
  let
38819
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   138
    val flip = Option.map not
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   139
    (* We include free variables, as well as constants, to handle locales. For
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   140
       each quantifiers that must necessarily be skolemized by the ATP, we
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   141
       introduce a fresh constant to simulate the effect of Skolemization. *)
38827
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   142
    fun do_const const (s, T) ts =
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   143
      add_pconst_to_table also_skolems (pconst thy const (s, T) ts)
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   144
      #> fold do_term ts
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   145
    and do_term t =
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   146
      case strip_comb t of
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   147
        (Const x, ts) => do_const true x ts
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   148
      | (Free x, ts) => do_const false x ts
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   149
      | (Abs (_, _, t'), ts) =>
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   150
        null ts ? add_pconst_to_table true (abs_name, [])
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   151
        #> fold do_term (t' :: ts)
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   152
      | (_, ts) => fold do_term ts
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   153
    fun do_quantifier will_surely_be_skolemized body_t =
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   154
      do_formula pos body_t
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   155
      #> (if also_skolems andalso will_surely_be_skolemized then
38823
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
   156
            add_pconst_to_table true (gensym skolem_prefix, [])
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   157
          else
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   158
            I)
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   159
    and do_term_or_formula T =
38692
89d3550d8e16 cosmetics
blanchet
parents: 38689
diff changeset
   160
      if is_formula_type T then do_formula NONE else do_term
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   161
    and do_formula pos t =
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   162
      case t of
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   163
        Const (@{const_name all}, _) $ Abs (_, _, body_t) =>
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   164
        do_quantifier (pos = SOME false) body_t
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   165
      | @{const "==>"} $ t1 $ t2 =>
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   166
        do_formula (flip pos) t1 #> do_formula pos t2
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   167
      | Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2 =>
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   168
        fold (do_term_or_formula T) [t1, t2]
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   169
      | @{const Trueprop} $ t1 => do_formula pos t1
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   170
      | @{const Not} $ t1 => do_formula (flip pos) t1
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   171
      | Const (@{const_name All}, _) $ Abs (_, _, body_t) =>
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   172
        do_quantifier (pos = SOME false) body_t
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   173
      | Const (@{const_name Ex}, _) $ Abs (_, _, body_t) =>
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   174
        do_quantifier (pos = SOME true) body_t
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   175
      | @{const HOL.conj} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   176
      | @{const HOL.disj} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
38786
e46e7a9cb622 formerly unnamed infix impliciation now named HOL.implies
haftmann
parents: 38752
diff changeset
   177
      | @{const HOL.implies} $ t1 $ t2 =>
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   178
        do_formula (flip pos) t1 #> do_formula pos t2
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   179
      | Const (@{const_name "op ="}, Type (_, [T, _])) $ t1 $ t2 =>
38587
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   180
        fold (do_term_or_formula T) [t1, t2]
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   181
      | Const (@{const_name If}, Type (_, [_, Type (_, [T, _])]))
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   182
        $ t1 $ t2 $ t3 =>
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   183
        do_formula NONE t1 #> fold (do_term_or_formula T) [t2, t3]
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   184
      | Const (@{const_name Ex1}, _) $ Abs (_, _, body_t) =>
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   185
        do_quantifier (is_some pos) body_t
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   186
      | Const (@{const_name Ball}, _) $ t1 $ Abs (_, _, body_t) =>
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   187
        do_quantifier (pos = SOME false)
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   188
                      (HOLogic.mk_imp (incr_boundvars 1 t1 $ Bound 0, body_t))
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   189
      | Const (@{const_name Bex}, _) $ t1 $ Abs (_, _, body_t) =>
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   190
        do_quantifier (pos = SOME true)
1317657d6aa9 fix the relevance filter so that it ignores If, Ex1, Ball, Bex
blanchet
parents: 38395
diff changeset
   191
                      (HOLogic.mk_conj (incr_boundvars 1 t1 $ Bound 0, body_t))
37537
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   192
      | (t0 as Const (_, @{typ bool})) $ t1 =>
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   193
        do_term t0 #> do_formula pos t1  (* theory constant *)
8e56d1ccf189 improve the new "natural formula" fact filter
blanchet
parents: 37515
diff changeset
   194
      | _ => do_term t
38819
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   195
  in Symtab.empty |> fold (do_formula pos) ts end
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   196
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   197
(*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
   198
  takes the given theory into account.*)
37616
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
   199
fun theory_const_prop_of theory_relevant th =
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   200
  if theory_relevant then
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   201
    let
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   202
      val name = Context.theory_name (theory_of_thm th)
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   203
      val t = Const (name ^ ". 1", @{typ bool})
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   204
    in t $ prop_of th end
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   205
  else
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   206
    prop_of th
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   207
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   208
(**** Constant / Type Frequencies ****)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   209
38743
69fa75354c58 simplify more code
blanchet
parents: 38742
diff changeset
   210
(* A two-dimensional symbol table counts frequencies of constants. It's keyed
69fa75354c58 simplify more code
blanchet
parents: 38742
diff changeset
   211
   first by constant name and second by its list of type instantiations. For the
38823
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
   212
   latter, we need a linear ordering on "pattern list". *)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   213
38823
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
   214
fun pattern_ord p =
38743
69fa75354c58 simplify more code
blanchet
parents: 38742
diff changeset
   215
  case p of
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   216
    (PVar, PVar) => EQUAL
38823
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
   217
  | (PVar, PApp _) => LESS
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
   218
  | (PApp _, PVar) => GREATER
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
   219
  | (PApp q1, PApp q2) =>
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
   220
    prod_ord fast_string_ord (dict_ord pattern_ord) (q1, q2)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   221
38743
69fa75354c58 simplify more code
blanchet
parents: 38742
diff changeset
   222
structure CTtab =
38823
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
   223
  Table(type key = pattern list val ord = dict_ord pattern_ord)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   224
38827
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   225
fun count_axiom_consts theory_relevant thy =
37503
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   226
  let
38827
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   227
    fun do_const const (s, T) ts =
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   228
      (* Two-dimensional table update. Constant maps to types maps to count. *)
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   229
      CTtab.map_default (pconst_args thy const (s, T) ts, 0) (Integer.add 1)
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   230
      |> Symtab.map_default (s, CTtab.empty)
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   231
      #> fold do_term ts
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   232
    and do_term t =
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   233
      case strip_comb t of
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   234
        (Const x, ts) => do_const true x ts
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   235
      | (Free x, ts) => do_const false x ts
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   236
      | (Abs (_, _, t'), ts) => fold do_term (t' :: ts)
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   237
      | (_, ts) => fold do_term ts
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   238
  in do_term o theory_const_prop_of theory_relevant o snd end
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   239
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   240
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   241
(**** Actual Filtering Code ****)
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   242
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   243
(*The frequency of a constant is the sum of those of all instances of its type.*)
38824
f74513bbe627 cosmetics
blanchet
parents: 38823
diff changeset
   244
fun pconst_freq match const_tab (c, ps) =
f74513bbe627 cosmetics
blanchet
parents: 38823
diff changeset
   245
  CTtab.fold (fn (qs, m) => match (ps, qs) ? Integer.add m)
38686
45eeee8d6b12 modified relevance filter
blanchet
parents: 38684
diff changeset
   246
             (the (Symtab.lookup const_tab c)) 0
45eeee8d6b12 modified relevance filter
blanchet
parents: 38684
diff changeset
   247
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   248
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   249
(* A surprising number of theorems contain only a few significant constants.
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   250
   These include all induction rules, and other general theorems. *)
37503
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   251
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   252
(* "log" seems best in practice. A constant function of one ignores the constant
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   253
   frequencies. *)
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   254
fun rel_log n = 1.0 + 2.0 / Math.ln (Real.fromInt n + 1.0)
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   255
fun irrel_log n = Math.ln (Real.fromInt n + 19.0) / 6.4
37503
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   256
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   257
(* FUDGE *)
38821
d0275b6c4e9d avoid needless "that" fact
blanchet
parents: 38820
diff changeset
   258
val abs_rel_weight = 0.5
d0275b6c4e9d avoid needless "that" fact
blanchet
parents: 38820
diff changeset
   259
val abs_irrel_weight = 2.0
d0275b6c4e9d avoid needless "that" fact
blanchet
parents: 38820
diff changeset
   260
val skolem_rel_weight = 2.0  (* impossible *)
d0275b6c4e9d avoid needless "that" fact
blanchet
parents: 38820
diff changeset
   261
val skolem_irrel_weight = 0.5
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   262
37503
c2dfa26b9da6 cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet
parents: 37502
diff changeset
   263
(* Computes a constant's weight, as determined by its frequency. *)
38821
d0275b6c4e9d avoid needless "that" fact
blanchet
parents: 38820
diff changeset
   264
fun generic_weight abs_weight skolem_weight logx f const_tab (c as (s, _)) =
38816
21a6f261595e lower penalty for Skolem constants
blanchet
parents: 38752
diff changeset
   265
  if s = abs_name then abs_weight
21a6f261595e lower penalty for Skolem constants
blanchet
parents: 38752
diff changeset
   266
  else if String.isPrefix skolem_prefix s then skolem_weight
38823
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
   267
  else logx (pconst_freq (match_patterns o f) const_tab c)
38821
d0275b6c4e9d avoid needless "that" fact
blanchet
parents: 38820
diff changeset
   268
d0275b6c4e9d avoid needless "that" fact
blanchet
parents: 38820
diff changeset
   269
val rel_weight = generic_weight abs_rel_weight skolem_rel_weight rel_log I
d0275b6c4e9d avoid needless "that" fact
blanchet
parents: 38820
diff changeset
   270
val irrel_weight = generic_weight abs_irrel_weight skolem_irrel_weight irrel_log
d0275b6c4e9d avoid needless "that" fact
blanchet
parents: 38820
diff changeset
   271
                                  swap
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   272
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   273
(* FUDGE *)
38889
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   274
fun locality_bonus General = 0.0
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   275
  | locality_bonus Theory = 0.5
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   276
  | locality_bonus Local = 1.0
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   277
  | locality_bonus Chained = 2.0
38751
01c4d14b2a61 add a bonus for chained facts, since they are likely to be relevant;
blanchet
parents: 38749
diff changeset
   278
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   279
fun axiom_weight loc const_tab relevant_consts axiom_consts =
38827
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   280
  case axiom_consts |> List.partition (pconst_hyper_mem I relevant_consts)
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   281
                    ||> filter_out (pconst_hyper_mem swap relevant_consts) of
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   282
    ([], _) => 0.0
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   283
  | (rel, irrel) =>
38889
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   284
    let
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   285
      val irrel = irrel |> filter_out (pconst_mem swap rel)
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   286
      val rel_weight = 0.0 |> fold (curry (op +) o rel_weight const_tab) rel
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   287
      val irrel_weight =
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   288
        ~ (locality_bonus loc)
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   289
        |> fold (curry (op +) o irrel_weight const_tab) irrel
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   290
      val res = rel_weight / (rel_weight + irrel_weight)
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   291
    in if Real.isFinite res then res else 0.0 end
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   292
38823
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
   293
fun pconsts_in_axiom thy t =
38825
4ec3cbd95f25 rename and simplify
blanchet
parents: 38824
diff changeset
   294
  Symtab.fold (fn (s, pss) => fold (cons o pair s) pss)
38823
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
   295
              (get_pconsts thy true (SOME true) [t]) []
38687
97509445c569 cache previous iteration's weights, and keep track of what's dirty and what's clean;
blanchet
parents: 38686
diff changeset
   296
fun pair_consts_axiom theory_relevant thy axiom =
38827
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   297
  case axiom |> snd |> theory_const_prop_of theory_relevant
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   298
             |> pconsts_in_axiom thy of
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   299
    [] => NONE
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   300
  | consts => SOME ((axiom, consts), NONE)
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   301
38699
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   302
type annotated_thm =
38823
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
   303
  (((unit -> string) * locality) * thm) * (string * pattern list) list
37505
d9af5c01dc4a added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet
parents: 37504
diff changeset
   304
38889
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   305
val max_imperfect_exp = 4.0
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   306
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   307
fun take_most_relevant max_max_imperfect max_relevant remaining_max
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   308
                       (candidates : (annotated_thm * real) list) =
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   309
  let
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   310
    val max_imperfect =
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   311
      Real.ceil (Math.pow (max_max_imperfect,
38889
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   312
                     Math.pow (Real.fromInt remaining_max
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   313
                               / Real.fromInt max_relevant, max_imperfect_exp)))
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   314
    val (perfect, imperfect) =
38889
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   315
      candidates |> sort (Real.compare o swap o pairself snd)
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   316
                 |> take_prefix (fn (_, w) => w > 0.99999)
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   317
    val ((accepts, more_rejects), rejects) =
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   318
      chop max_imperfect imperfect |>> append perfect |>> chop remaining_max
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   319
  in
38889
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   320
    trace_msg (fn () =>
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   321
        "Actually passed (" ^ Int.toString (length accepts) ^ " of " ^
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   322
        Int.toString (length candidates) ^ "): " ^
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   323
        (accepts |> map (fn ((((name, _), _), _), weight) =>
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   324
                            name () ^ " [" ^ Real.toString weight ^ "]")
38745
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   325
                 |> commas));
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   326
    (accepts, more_rejects @ rejects)
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   327
  end
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   328
38819
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   329
fun if_empty_replace_with_locality thy axioms loc tab =
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   330
  if Symtab.is_empty tab then
38823
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
   331
    get_pconsts thy false (SOME false)
38819
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   332
        (map_filter (fn ((_, loc'), th) =>
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   333
                        if loc' = loc then SOME (prop_of th) else NONE) axioms)
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   334
  else
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   335
    tab
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   336
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   337
(* FUDGE *)
38683
23266607cb81 if no facts were selected on first iteration, try again with a lower threshold
blanchet
parents: 38682
diff changeset
   338
val threshold_divisor = 2.0
23266607cb81 if no facts were selected on first iteration, try again with a lower threshold
blanchet
parents: 38682
diff changeset
   339
val ridiculous_threshold = 0.1
38889
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   340
val max_max_imperfect_fudge_factor = 0.5
38683
23266607cb81 if no facts were selected on first iteration, try again with a lower threshold
blanchet
parents: 38682
diff changeset
   341
38745
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   342
fun relevance_filter ctxt threshold0 decay max_relevant theory_relevant
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   343
                     ({add, del, ...} : relevance_override) axioms goal_ts =
38739
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   344
  let
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   345
    val thy = ProofContext.theory_of ctxt
38819
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   346
    val const_tab =
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   347
      fold (count_axiom_consts theory_relevant thy) axioms Symtab.empty
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   348
    val goal_const_tab =
38823
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
   349
      get_pconsts thy false (SOME false) goal_ts
38819
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   350
      |> fold (if_empty_replace_with_locality thy axioms)
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   351
              [Chained, Local, Theory]
38739
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   352
    val add_thms = maps (ProofContext.get_fact ctxt) add
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   353
    val del_thms = maps (ProofContext.get_fact ctxt) del
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   354
    val max_max_imperfect =
38889
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   355
      Math.sqrt (Real.fromInt max_relevant) * max_max_imperfect_fudge_factor
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   356
    fun iter j remaining_max threshold rel_const_tab hopeless hopeful =
38739
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   357
      let
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   358
        fun game_over rejects =
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   359
          (* Add "add:" facts. *)
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   360
          if null add_thms then
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   361
            []
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   362
          else
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   363
            map_filter (fn ((p as (_, th), _), _) =>
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   364
                           if member Thm.eq_thm add_thms th then SOME p
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   365
                           else NONE) rejects
38889
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   366
        fun relevant [] rejects [] =
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   367
            (* Nothing has been added this iteration. *)
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   368
            if j = 0 andalso threshold >= ridiculous_threshold then
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   369
              (* First iteration? Try again. *)
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   370
              iter 0 max_relevant (threshold / threshold_divisor) rel_const_tab
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   371
                   hopeless hopeful
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   372
            else
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   373
              game_over (rejects @ hopeless)
38889
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   374
          | relevant candidates rejects [] =
38739
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   375
            let
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   376
              val (accepts, more_rejects) =
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   377
                take_most_relevant max_max_imperfect max_relevant remaining_max
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   378
                                   candidates
38739
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   379
              val rel_const_tab' =
38745
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   380
                rel_const_tab
38823
828e68441a2f renaming + treat "TFree" better in "pattern_for_type"
blanchet
parents: 38822
diff changeset
   381
                |> fold (add_pconst_to_table false)
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   382
                        (maps (snd o fst) accepts)
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   383
              fun is_dirty (c, _) =
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   384
                Symtab.lookup rel_const_tab' c <> Symtab.lookup rel_const_tab c
38745
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   385
              val (hopeful_rejects, hopeless_rejects) =
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   386
                 (rejects @ hopeless, ([], []))
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   387
                 |-> fold (fn (ax as (_, consts), old_weight) =>
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   388
                              if exists is_dirty consts then
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   389
                                apfst (cons (ax, NONE))
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   390
                              else
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   391
                                apsnd (cons (ax, old_weight)))
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   392
                 |>> append (more_rejects
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   393
                             |> map (fn (ax as (_, consts), old_weight) =>
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   394
                                        (ax, if exists is_dirty consts then NONE
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   395
                                             else SOME old_weight)))
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   396
              val threshold =
38822
aa0101e618e2 fix threshold computation + remove "op =" from relevant constants
blanchet
parents: 38821
diff changeset
   397
                1.0 - (1.0 - threshold)
aa0101e618e2 fix threshold computation + remove "op =" from relevant constants
blanchet
parents: 38821
diff changeset
   398
                      * Math.pow (decay, Real.fromInt (length accepts))
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   399
              val remaining_max = remaining_max - length accepts
38739
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   400
            in
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   401
              trace_msg (fn () => "New or updated constants: " ^
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   402
                  commas (rel_const_tab' |> Symtab.dest
38822
aa0101e618e2 fix threshold computation + remove "op =" from relevant constants
blanchet
parents: 38821
diff changeset
   403
                          |> subtract (op =) (rel_const_tab |> Symtab.dest)
38827
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   404
                          |> map string_for_hyper_pconst));
38745
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   405
              map (fst o fst) accepts @
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   406
              (if remaining_max = 0 then
38745
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   407
                 game_over (hopeful_rejects @ map (apsnd SOME) hopeless_rejects)
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   408
               else
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   409
                 iter (j + 1) remaining_max threshold rel_const_tab'
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   410
                      hopeless_rejects hopeful_rejects)
38739
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   411
            end
38889
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   412
          | relevant candidates rejects
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   413
                     (((ax as (((_, loc), th), axiom_consts)), cached_weight)
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   414
                      :: hopeful) =
38739
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   415
            let
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   416
              val weight =
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   417
                case cached_weight of
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   418
                  SOME w => w
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   419
                | NONE => axiom_weight loc const_tab rel_const_tab axiom_consts
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   420
(* TODO: experiment
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   421
val name = fst (fst (fst ax)) ()
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   422
val _ = if String.isPrefix "lift.simps(3" name then
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   423
tracing ("*** " ^ name ^ PolyML.makestring (debug_axiom_weight const_tab rel_const_tab axiom_consts))
38747
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   424
else
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   425
()
b264ae66cede fiddle with relevance filter
blanchet
parents: 38745
diff changeset
   426
*)
38739
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   427
            in
38741
7635bf8918a1 get rid of "defs_relevant" feature;
blanchet
parents: 38739
diff changeset
   428
              if weight >= threshold then
38889
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   429
                relevant ((ax, weight) :: candidates) rejects hopeful
38739
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   430
              else
38889
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   431
                relevant candidates ((ax, weight) :: rejects) hopeful
38739
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   432
            end
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   433
        in
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   434
          trace_msg (fn () =>
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   435
              "ITERATION " ^ string_of_int j ^ ": current threshold: " ^
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   436
              Real.toString threshold ^ ", constants: " ^
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   437
              commas (rel_const_tab |> Symtab.dest
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   438
                      |> filter (curry (op <>) [] o snd)
38827
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   439
                      |> map string_for_hyper_pconst));
38889
d0e3f68dde63 fiddle with the relevance filter
blanchet
parents: 38829
diff changeset
   440
          relevant [] [] hopeful
38739
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   441
        end
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   442
  in
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   443
    axioms |> filter_out (member Thm.eq_thm del_thms o snd)
38827
cf01645cbbce extended relevance filter with first-order term matching
blanchet
parents: 38825
diff changeset
   444
           |> map_filter (pair_consts_axiom theory_relevant thy)
38819
71c9f61516cd if the goal contains no constants or frees, fall back on chained facts, then on local facts, etc., instead of generating a trivial ATP problem
blanchet
parents: 38818
diff changeset
   445
           |> iter 0 max_relevant threshold0 goal_const_tab []
38739
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   446
           |> tap (fn res => trace_msg (fn () =>
38686
45eeee8d6b12 modified relevance filter
blanchet
parents: 38684
diff changeset
   447
                                "Total relevant: " ^ Int.toString (length res)))
38739
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   448
  end
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   449
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   450
24287
c857dac06da6 combining the relevance filter with res_atp
paulson
parents: 24286
diff changeset
   451
(***************************************************************)
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   452
(* Retrieving and filtering lemmas                             *)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   453
(***************************************************************)
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   454
33022
c95102496490 Removal of the unused atpset concept, the atp attribute and some related code.
paulson
parents: 32994
diff changeset
   455
(*** retrieve lemmas and filter them ***)
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   456
20757
fe84fe0dfd30 Definitions produced by packages are now blacklisted.
paulson
parents: 20661
diff changeset
   457
(*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
   458
  "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
   459
fun is_package_def a =
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30291
diff changeset
   460
  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
   461
  in
552d20ff9a95 Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents: 21588
diff changeset
   462
     length names > 2 andalso
552d20ff9a95 Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents: 21588
diff changeset
   463
     not (hd names = "local") andalso
552d20ff9a95 Removal of theorem tagging, which the ATP linkup no longer requires.
paulson
parents: 21588
diff changeset
   464
     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
   465
  end;
20757
fe84fe0dfd30 Definitions produced by packages are now blacklisted.
paulson
parents: 20661
diff changeset
   466
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   467
fun make_fact_table xs =
37616
c8d2d84d6011 always perform relevance filtering on original formulas
blanchet
parents: 37580
diff changeset
   468
  fold (Termtab.update o `(prop_of o snd)) xs Termtab.empty
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   469
fun make_unique xs = Termtab.fold (cons o snd) (make_fact_table xs) []
19768
9afd9b9c47d0 ATP/res_clasimpset.ML has been merged into res_atp.ML.
mengj
parents: 19746
diff changeset
   470
37626
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   471
(* FIXME: put other record thms here, or declare as "no_atp" *)
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   472
val multi_base_blacklist =
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   473
  ["defs", "select_defs", "update_defs", "induct", "inducts", "split", "splits",
38682
3a203da3f89b weed out junk in relevance filter
blanchet
parents: 38681
diff changeset
   474
   "split_asm", "cases", "ext_cases", "eq.simps", "eq.refl", "nchotomy",
3a203da3f89b weed out junk in relevance filter
blanchet
parents: 38681
diff changeset
   475
   "case_cong", "weak_case_cong"]
3a203da3f89b weed out junk in relevance filter
blanchet
parents: 38681
diff changeset
   476
  |> map (prefix ".")
37626
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   477
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   478
val max_lambda_nesting = 3
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   479
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   480
fun term_has_too_many_lambdas max (t1 $ t2) =
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   481
    exists (term_has_too_many_lambdas max) [t1, t2]
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   482
  | term_has_too_many_lambdas max (Abs (_, _, t)) =
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   483
    max = 0 orelse term_has_too_many_lambdas (max - 1) t
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   484
  | term_has_too_many_lambdas _ _ = false
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   485
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   486
(* Don't count nested lambdas at the level of formulas, since they are
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   487
   quantifiers. *)
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   488
fun formula_has_too_many_lambdas Ts (Abs (_, T, t)) =
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   489
    formula_has_too_many_lambdas (T :: Ts) t
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   490
  | formula_has_too_many_lambdas Ts t =
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   491
    if is_formula_type (fastype_of1 (Ts, t)) then
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   492
      exists (formula_has_too_many_lambdas Ts) (#2 (strip_comb t))
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   493
    else
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   494
      term_has_too_many_lambdas max_lambda_nesting t
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   495
38692
89d3550d8e16 cosmetics
blanchet
parents: 38689
diff changeset
   496
(* The max apply depth of any "metis" call in "Metis_Examples" (on 2007-10-31)
37626
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   497
   was 11. *)
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   498
val max_apply_depth = 15
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   499
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   500
fun apply_depth (f $ t) = Int.max (apply_depth f, apply_depth t + 1)
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   501
  | apply_depth (Abs (_, _, t)) = apply_depth t
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   502
  | apply_depth _ = 0
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   503
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   504
fun is_formula_too_complex t =
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   505
  apply_depth t > max_apply_depth orelse formula_has_too_many_lambdas [] t
37626
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   506
37543
2e733b0a963c a76ace919f1c wasn't quite right; second try
blanchet
parents: 37538
diff changeset
   507
val exists_sledgehammer_const =
37626
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   508
  exists_Const (fn (s, _) => String.isPrefix sledgehammer_prefix s)
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   509
38821
d0275b6c4e9d avoid needless "that" fact
blanchet
parents: 38820
diff changeset
   510
fun is_metastrange_theorem th =
37626
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   511
  case head_of (concl_of th) of
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   512
      Const (a, _) => (a <> @{const_name Trueprop} andalso
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   513
                       a <> @{const_name "=="})
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   514
    | _ => false
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   515
38821
d0275b6c4e9d avoid needless "that" fact
blanchet
parents: 38820
diff changeset
   516
fun is_that_fact th =
d0275b6c4e9d avoid needless "that" fact
blanchet
parents: 38820
diff changeset
   517
  String.isSuffix (Long_Name.separator ^ Obtain.thatN) (Thm.get_name_hint th)
d0275b6c4e9d avoid needless "that" fact
blanchet
parents: 38820
diff changeset
   518
  andalso exists_subterm (fn Free (s, _) => s = Name.skolem Auto_Bind.thesisN
d0275b6c4e9d avoid needless "that" fact
blanchet
parents: 38820
diff changeset
   519
                           | _ => false) (prop_of th)
d0275b6c4e9d avoid needless "that" fact
blanchet
parents: 38820
diff changeset
   520
37626
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   521
val type_has_top_sort =
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   522
  exists_subtype (fn TFree (_, []) => true | TVar (_, []) => true | _ => false)
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37616
diff changeset
   523
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   524
(**** Predicates to detect unwanted facts (prolific or likely to cause
37347
635425a442e8 show more respect for user-specified facts, even if they could lead to unsound proofs + don't throw away "unsound" theorems in "full_type" mode, since they are then sound
blanchet
parents: 37345
diff changeset
   525
      unsoundness) ****)
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   526
38289
74dd8dd33512 adapt "too_general_equality" blacklisting to the new FOF context, where quantifiers are sometimes present
blanchet
parents: 38279
diff changeset
   527
(* Too general means, positive equality literal with a variable X as one
74dd8dd33512 adapt "too_general_equality" blacklisting to the new FOF context, where quantifiers are sometimes present
blanchet
parents: 38279
diff changeset
   528
   operand, when X does not occur properly in the other operand. This rules out
74dd8dd33512 adapt "too_general_equality" blacklisting to the new FOF context, where quantifiers are sometimes present
blanchet
parents: 38279
diff changeset
   529
   clearly inconsistent facts such as X = a | X = b, though it by no means
74dd8dd33512 adapt "too_general_equality" blacklisting to the new FOF context, where quantifiers are sometimes present
blanchet
parents: 38279
diff changeset
   530
   guarantees soundness. *)
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   531
38289
74dd8dd33512 adapt "too_general_equality" blacklisting to the new FOF context, where quantifiers are sometimes present
blanchet
parents: 38279
diff changeset
   532
(* Unwanted equalities are those between a (bound or schematic) variable that
74dd8dd33512 adapt "too_general_equality" blacklisting to the new FOF context, where quantifiers are sometimes present
blanchet
parents: 38279
diff changeset
   533
   does not properly occur in the second operand. *)
38607
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   534
val is_exhaustive_finite =
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   535
  let
38629
3387432c18af more work on finite axiom detection
blanchet
parents: 38627
diff changeset
   536
    fun is_bad_equal (Var z) t =
3387432c18af more work on finite axiom detection
blanchet
parents: 38627
diff changeset
   537
        not (exists_subterm (fn Var z' => z = z' | _ => false) t)
3387432c18af more work on finite axiom detection
blanchet
parents: 38627
diff changeset
   538
      | is_bad_equal (Bound j) t = not (loose_bvar1 (t, j))
3387432c18af more work on finite axiom detection
blanchet
parents: 38627
diff changeset
   539
      | is_bad_equal _ _ = false
3387432c18af more work on finite axiom detection
blanchet
parents: 38627
diff changeset
   540
    fun do_equals t1 t2 = is_bad_equal t1 t2 orelse is_bad_equal t2 t1
38607
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   541
    fun do_formula pos t =
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   542
      case (pos, t) of
38615
4e1d828ee514 improve "x = A | x = B | x = C"-style axiom detection
blanchet
parents: 38611
diff changeset
   543
        (_, @{const Trueprop} $ t1) => do_formula pos t1
38607
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   544
      | (true, Const (@{const_name all}, _) $ Abs (_, _, t')) =>
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   545
        do_formula pos t'
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   546
      | (true, Const (@{const_name All}, _) $ Abs (_, _, t')) =>
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   547
        do_formula pos t'
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   548
      | (false, Const (@{const_name Ex}, _) $ Abs (_, _, t')) =>
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   549
        do_formula pos t'
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   550
      | (_, @{const "==>"} $ t1 $ t2) =>
38629
3387432c18af more work on finite axiom detection
blanchet
parents: 38627
diff changeset
   551
        do_formula (not pos) t1 andalso
3387432c18af more work on finite axiom detection
blanchet
parents: 38627
diff changeset
   552
        (t2 = @{prop False} orelse do_formula pos t2)
38786
e46e7a9cb622 formerly unnamed infix impliciation now named HOL.implies
haftmann
parents: 38752
diff changeset
   553
      | (_, @{const HOL.implies} $ t1 $ t2) =>
38629
3387432c18af more work on finite axiom detection
blanchet
parents: 38627
diff changeset
   554
        do_formula (not pos) t1 andalso
3387432c18af more work on finite axiom detection
blanchet
parents: 38627
diff changeset
   555
        (t2 = @{const False} orelse do_formula pos t2)
38607
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   556
      | (_, @{const Not} $ t1) => do_formula (not pos) t1
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   557
      | (true, @{const HOL.disj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   558
      | (false, @{const HOL.conj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
38607
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   559
      | (true, Const (@{const_name "op ="}, _) $ t1 $ t2) => do_equals t1 t2
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   560
      | (true, Const (@{const_name "=="}, _) $ t1 $ t2) => do_equals t1 t2
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   561
      | _ => false
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   562
  in do_formula true end
a2abe8c2a1c2 generalize the "too general equality" code to handle facts like "x ~= A ==> x = B"
blanchet
parents: 38606
diff changeset
   563
38592
ae6bb801e583 bound variables can be just as evil as schematic variables and lead to unsound proofs (e.g. "all_bool_eq")
blanchet
parents: 38587
diff changeset
   564
fun has_bound_or_var_of_type tycons =
ae6bb801e583 bound variables can be just as evil as schematic variables and lead to unsound proofs (e.g. "all_bool_eq")
blanchet
parents: 38587
diff changeset
   565
  exists_subterm (fn Var (_, Type (s, _)) => member (op =) tycons s
ae6bb801e583 bound variables can be just as evil as schematic variables and lead to unsound proofs (e.g. "all_bool_eq")
blanchet
parents: 38587
diff changeset
   566
                   | Abs (_, Type (s, _), _) => member (op =) tycons s
ae6bb801e583 bound variables can be just as evil as schematic variables and lead to unsound proofs (e.g. "all_bool_eq")
blanchet
parents: 38587
diff changeset
   567
                   | _ => false)
21431
ef9080e7dbbc Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
paulson
parents: 21397
diff changeset
   568
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   569
(* Facts are forbidden to contain variables of these types. The typical reason
37347
635425a442e8 show more respect for user-specified facts, even if they could lead to unsound proofs + don't throw away "unsound" theorems in "full_type" mode, since they are then sound
blanchet
parents: 37345
diff changeset
   570
   is that they lead to unsoundness. Note that "unit" satisfies numerous
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   571
   equations like "?x = ()". The resulting clauses will have no type constraint,
37347
635425a442e8 show more respect for user-specified facts, even if they could lead to unsound proofs + don't throw away "unsound" theorems in "full_type" mode, since they are then sound
blanchet
parents: 37345
diff changeset
   572
   yielding false proofs. Even "bool" leads to many unsound proofs, though only
635425a442e8 show more respect for user-specified facts, even if they could lead to unsound proofs + don't throw away "unsound" theorems in "full_type" mode, since they are then sound
blanchet
parents: 37345
diff changeset
   573
   for higher-order problems. *)
38592
ae6bb801e583 bound variables can be just as evil as schematic variables and lead to unsound proofs (e.g. "all_bool_eq")
blanchet
parents: 38587
diff changeset
   574
val dangerous_types = [@{type_name unit}, @{type_name bool}, @{type_name prop}];
22217
a5d983f7113f Tidying; more debugging information. New reference unwanted_types.
paulson
parents: 22193
diff changeset
   575
38085
cc44e887246c avoid "clause" and "cnf" terminology where it no longer makes sense
blanchet
parents: 38027
diff changeset
   576
(* Facts containing variables of type "unit" or "bool" or of the form
38290
581a402a80f0 prevent ATP thread for staying around for 1 minute if an exception occurred earlier;
blanchet
parents: 38289
diff changeset
   577
   "ALL x. x = A | x = B | x = C" are likely to lead to unsound proofs if types
581a402a80f0 prevent ATP thread for staying around for 1 minute if an exception occurred earlier;
blanchet
parents: 38289
diff changeset
   578
   are omitted. *)
38593
85aef7587cf1 minor cleanup
blanchet
parents: 38592
diff changeset
   579
fun is_dangerous_term full_types t =
38609
6220e5ab32f7 more fiddling with Sledgehammer's "add:" option
blanchet
parents: 38607
diff changeset
   580
  not full_types andalso
38679
2cfd0777580f destroy elim rules before checking for finite exhaustive facts
blanchet
parents: 38652
diff changeset
   581
  let val t = transform_elim_term t in
2cfd0777580f destroy elim rules before checking for finite exhaustive facts
blanchet
parents: 38652
diff changeset
   582
    has_bound_or_var_of_type dangerous_types t orelse
2cfd0777580f destroy elim rules before checking for finite exhaustive facts
blanchet
parents: 38652
diff changeset
   583
    is_exhaustive_finite t
2cfd0777580f destroy elim rules before checking for finite exhaustive facts
blanchet
parents: 38652
diff changeset
   584
  end
21470
7c1b59ddcd56 Consolidation of code to "blacklist" unhelpful theorems, including record
paulson
parents: 21431
diff changeset
   585
38627
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   586
fun is_theorem_bad_for_atps full_types thm =
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   587
  let val t = prop_of thm in
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   588
    is_formula_too_complex t orelse exists_type type_has_top_sort t orelse
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   589
    is_dangerous_term full_types t orelse exists_sledgehammer_const t orelse
38821
d0275b6c4e9d avoid needless "that" fact
blanchet
parents: 38820
diff changeset
   590
    is_metastrange_theorem thm orelse is_that_fact thm
38627
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   591
  end
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   592
38696
4c6b65d6a135 quote facts whose names collide with a keyword or command name (cf. "subclass" in "Jinja/J/TypeSafe.thy")
blanchet
parents: 38692
diff changeset
   593
fun all_name_thms_pairs ctxt reserved full_types add_thms chained_ths =
38627
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   594
  let
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   595
    val thy = ProofContext.theory_of ctxt
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   596
    val thy_prefix = Context.theory_name thy ^ Long_Name.separator
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   597
    val global_facts = PureThy.facts_of thy
38644
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   598
    val local_facts = ProofContext.facts_of ctxt
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   599
    val named_locals = local_facts |> Facts.dest_static []
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   600
    val is_chained = member Thm.eq_thm chained_ths
38818
61cf050f8b2e improve SPASS hack, when a clause comes from several facts
blanchet
parents: 38816
diff changeset
   601
    (* Unnamed nonchained formulas with schematic variables are omitted, because
61cf050f8b2e improve SPASS hack, when a clause comes from several facts
blanchet
parents: 38816
diff changeset
   602
       they are rejected by the backticks (`...`) parser for some reason. *)
38738
0ce517c1970f make sure that "undo_ascii_of" is the inverse of "ascii_of", also for non-printable characters -- and avoid those in ``-style facts
blanchet
parents: 38699
diff changeset
   603
    fun is_good_unnamed_local th =
38820
d0f98bd81a85 add nameless chained facts to the pool of things known to Sledgehammer
blanchet
parents: 38819
diff changeset
   604
      not (Thm.has_name_hint th) andalso
d0f98bd81a85 add nameless chained facts to the pool of things known to Sledgehammer
blanchet
parents: 38819
diff changeset
   605
      (not (exists_subterm is_Var (prop_of th)) orelse (is_chained th)) andalso
38738
0ce517c1970f make sure that "undo_ascii_of" is the inverse of "ascii_of", also for non-printable characters -- and avoid those in ``-style facts
blanchet
parents: 38699
diff changeset
   606
      forall (fn (_, ths) => not (member Thm.eq_thm ths th)) named_locals
38644
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   607
    val unnamed_locals =
38820
d0f98bd81a85 add nameless chained facts to the pool of things known to Sledgehammer
blanchet
parents: 38819
diff changeset
   608
      union Thm.eq_thm (Facts.props local_facts) chained_ths
d0f98bd81a85 add nameless chained facts to the pool of things known to Sledgehammer
blanchet
parents: 38819
diff changeset
   609
      |> filter is_good_unnamed_local |> map (pair "" o single)
38627
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   610
    val full_space =
38738
0ce517c1970f make sure that "undo_ascii_of" is the inverse of "ascii_of", also for non-printable characters -- and avoid those in ``-style facts
blanchet
parents: 38699
diff changeset
   611
      Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts)
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   612
    fun add_facts global foldx facts =
38699
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   613
      foldx (fn (name0, ths) =>
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   614
        if name0 <> "" andalso
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   615
           forall (not o member Thm.eq_thm add_thms) ths andalso
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   616
           (Facts.is_concealed facts name0 orelse
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   617
            (respect_no_atp andalso is_package_def name0) orelse
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   618
            exists (fn s => String.isSuffix s name0) multi_base_blacklist orelse
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   619
            String.isSuffix "_def_raw" (* FIXME: crude hack *) name0) then
38627
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   620
          I
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   621
        else
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   622
          let
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   623
            val base_loc =
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   624
              if not global then Local
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   625
              else if String.isPrefix thy_prefix name0 then Theory
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   626
              else General
38699
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   627
            val multi = length ths > 1
38696
4c6b65d6a135 quote facts whose names collide with a keyword or command name (cf. "subclass" in "Jinja/J/TypeSafe.thy")
blanchet
parents: 38692
diff changeset
   628
            fun backquotify th =
4c6b65d6a135 quote facts whose names collide with a keyword or command name (cf. "subclass" in "Jinja/J/TypeSafe.thy")
blanchet
parents: 38692
diff changeset
   629
              "`" ^ Print_Mode.setmp [Print_Mode.input]
4c6b65d6a135 quote facts whose names collide with a keyword or command name (cf. "subclass" in "Jinja/J/TypeSafe.thy")
blanchet
parents: 38692
diff changeset
   630
                                 (Syntax.string_of_term ctxt) (prop_of th) ^ "`"
38738
0ce517c1970f make sure that "undo_ascii_of" is the inverse of "ascii_of", also for non-printable characters -- and avoid those in ``-style facts
blanchet
parents: 38699
diff changeset
   631
              |> String.translate (fn c => if Char.isPrint c then str c else "")
0ce517c1970f make sure that "undo_ascii_of" is the inverse of "ascii_of", also for non-printable characters -- and avoid those in ``-style facts
blanchet
parents: 38699
diff changeset
   632
              |> simplify_spaces
38699
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   633
            fun check_thms a =
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   634
              case try (ProofContext.get_thms ctxt) a of
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   635
                NONE => false
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   636
              | SOME ths' => Thm.eq_thms (ths, ths')
38627
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   637
          in
38699
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   638
            pair 1
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   639
            #> fold (fn th => fn (j, rest) =>
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   640
                 (j + 1,
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   641
                  if is_theorem_bad_for_atps full_types th andalso
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   642
                     not (member Thm.eq_thm add_thms th) then
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   643
                    rest
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   644
                  else
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   645
                    (((fn () =>
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   646
                          if name0 = "" then
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   647
                            th |> backquotify
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   648
                          else
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   649
                            let
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   650
                              val name1 = Facts.extern facts name0
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   651
                              val name2 = Name_Space.extern full_space name0
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   652
                            in
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   653
                              case find_first check_thms [name1, name2, name0] of
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   654
                                SOME name => repair_name reserved multi j name
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   655
                              | NONE => ""
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   656
                            end), if is_chained th then Chained else base_loc),
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   657
                      (multi, th)) :: rest)) ths
38699
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   658
            #> snd
38627
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   659
          end)
38644
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   660
  in
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   661
    [] |> add_facts false fold local_facts (unnamed_locals @ named_locals)
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   662
       |> add_facts true Facts.fold_static global_facts global_facts
38644
25bbbaf7ce65 don't penalize abstractions in relevance filter + support nameless `foo`-style facts
blanchet
parents: 38629
diff changeset
   663
  end
38627
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   664
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   665
(* The single-name theorems go after the multiple-name ones, so that single
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   666
   names are preferred when both are available. *)
38699
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   667
fun name_thm_pairs ctxt respect_no_atp =
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   668
  List.partition (fst o snd) #> op @ #> map (apsnd snd)
38699
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   669
  #> respect_no_atp ? filter_out (No_ATPs.member ctxt o snd)
38627
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   670
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   671
(***************************************************************)
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   672
(* ATP invocation methods setup                                *)
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   673
(***************************************************************)
760a2d5cc671 make sure minimizer facts go through "transform_elim_theorems"
blanchet
parents: 38617
diff changeset
   674
38745
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   675
fun relevant_facts full_types (threshold0, threshold1) max_relevant
38744
2b6333f78a9e make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents: 38743
diff changeset
   676
                   theory_relevant (relevance_override as {add, del, only})
37995
06f02b15ef8a generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents: 37626
diff changeset
   677
                   (ctxt, (chained_ths, _)) hyp_ts concl_t =
37538
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   678
  let
38822
aa0101e618e2 fix threshold computation + remove "op =" from relevant constants
blanchet
parents: 38821
diff changeset
   679
    val decay = Math.pow ((1.0 - threshold1) / (1.0 - threshold0),
aa0101e618e2 fix threshold computation + remove "op =" from relevant constants
blanchet
parents: 38821
diff changeset
   680
                          1.0 / Real.fromInt (max_relevant + 1))
37538
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   681
    val add_thms = maps (ProofContext.get_fact ctxt) add
38696
4c6b65d6a135 quote facts whose names collide with a keyword or command name (cf. "subclass" in "Jinja/J/TypeSafe.thy")
blanchet
parents: 38692
diff changeset
   682
    val reserved = reserved_isar_keyword_table ()
37538
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   683
    val axioms =
38699
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   684
      (if only then
38752
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   685
         maps (map (fn ((name, loc), th) => ((K name, loc), (true, th)))
6628adcae4a7 consider "locality" when assigning weights to facts
blanchet
parents: 38751
diff changeset
   686
               o name_thm_pairs_from_ref ctxt reserved chained_ths) add
38699
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   687
       else
27378b4a776b compute names lazily;
blanchet
parents: 38698
diff changeset
   688
         all_name_thms_pairs ctxt reserved full_types add_thms chained_ths)
38688
b2ae937a134b optimize fact filter some more
blanchet
parents: 38687
diff changeset
   689
      |> name_thm_pairs ctxt (respect_no_atp andalso not only)
38595
bbb0982656eb make sure that "add:" doesn't influence the relevance filter too much
blanchet
parents: 38594
diff changeset
   690
      |> make_unique
37538
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   691
  in
38688
b2ae937a134b optimize fact filter some more
blanchet
parents: 38687
diff changeset
   692
    trace_msg (fn () => "Considering " ^ Int.toString (length axioms) ^
b2ae937a134b optimize fact filter some more
blanchet
parents: 38687
diff changeset
   693
                        " theorems");
38745
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   694
    (if threshold0 > 1.0 orelse threshold0 > threshold1 then
38739
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   695
       []
38745
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   696
     else if threshold0 < 0.0 then
38739
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   697
       axioms
8b8ed80b5699 renamed "relevance_convergence" to "relevance_decay"
blanchet
parents: 38738
diff changeset
   698
     else
38745
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   699
       relevance_filter ctxt threshold0 decay max_relevant theory_relevant
ad577fd62ee4 reorganize options regarding to the relevance threshold and decay
blanchet
parents: 38744
diff changeset
   700
                        relevance_override axioms (concl_t :: hyp_ts))
38822
aa0101e618e2 fix threshold computation + remove "op =" from relevant constants
blanchet
parents: 38821
diff changeset
   701
    |> map (apfst (apfst (fn f => f ())))
37538
97ab019d5ac8 make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet
parents: 37537
diff changeset
   702
  end
30536
07b4f050e4df split relevance-filter and writing of problem-files;
immler@in.tum.de
parents: 30364
diff changeset
   703
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   704
end;