src/HOL/Tools/Sledgehammer/clausifier.ML
author blanchet
Fri, 25 Jun 2010 16:15:03 +0200
changeset 37574 b8c1f4c46983
parent 37572 src/HOL/Tools/Sledgehammer/sledgehammer_fact_preprocessor.ML@a899f9506f39
child 37584 2e289dc13615
permissions -rw-r--r--
renamed "Sledgehammer_Fact_Preprocessor" to "Clausifier"; the new name reflects that it's not used only by Sledgehammer (but also by "meson" and "metis") and that it doesn't only clausify facts (but also goals)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37574
b8c1f4c46983 renamed "Sledgehammer_Fact_Preprocessor" to "Clausifier";
blanchet
parents: 37572
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/clausifier.ML
33311
49cd8abb2863 proper header;
wenzelm
parents: 33306
diff changeset
     2
    Author:     Jia Meng, Cambridge University Computer Laboratory
36393
be73a2b2443b support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents: 36228
diff changeset
     3
    Author:     Jasmin Blanchette, TU Muenchen
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     4
20461
d689ad772b2c skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents: 20457
diff changeset
     5
Transformation of axiom rules (elim/intro/etc) into CNF forms.
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     6
*)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     7
37574
b8c1f4c46983 renamed "Sledgehammer_Fact_Preprocessor" to "Clausifier";
blanchet
parents: 37572
diff changeset
     8
signature CLAUSIFIER =
21505
13d4dba99337 prefer Proof.context over Context.generic;
wenzelm
parents: 21470
diff changeset
     9
sig
37500
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
    10
  type cnf_thm = thm * ((string * int) * thm)
37543
2e733b0a963c a76ace919f1c wasn't quite right; second try
blanchet
parents: 37541
diff changeset
    11
2e733b0a963c a76ace919f1c wasn't quite right; second try
blanchet
parents: 37541
diff changeset
    12
  val sledgehammer_prefix: string
37171
fc1e20373e6a make sure chained facts appear in Isar proofs generated by Sledgehammer -- otherwise the proof won't work
blanchet
parents: 36945
diff changeset
    13
  val chained_prefix: string
32955
4a78daeb012b local channels for tracing/debugging;
wenzelm
parents: 32740
diff changeset
    14
  val trace: bool Unsynchronized.ref
4a78daeb012b local channels for tracing/debugging;
wenzelm
parents: 32740
diff changeset
    15
  val trace_msg: (unit -> string) -> unit
37567
02e4ccd512b6 further reduce dependencies on "sledgehammer_fact_filter.ML"
blanchet
parents: 37543
diff changeset
    16
  val name_thms_pair_from_ref :
02e4ccd512b6 further reduce dependencies on "sledgehammer_fact_filter.ML"
blanchet
parents: 37543
diff changeset
    17
    Proof.context -> thm list -> Facts.ref -> string * thm list
37410
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37403
diff changeset
    18
  val skolem_theory_name: string
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
    19
  val skolem_prefix: string
36492
60532b9bcd1c save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
blanchet
parents: 36478
diff changeset
    20
  val skolem_infix: string
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
    21
  val is_skolem_const_name: string -> bool
37572
a899f9506f39 more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents: 37571
diff changeset
    22
  val num_type_args: theory -> string -> int
27179
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
    23
  val cnf_axiom: theory -> thm -> thm list
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
    24
  val multi_base_blacklist: string list
37348
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
    25
  val is_theorem_bad_for_atps: thm -> bool
35568
8fbbfc39508f renamed type_has_empty_sort to type_has_topsort -- {} is the full universal sort;
wenzelm
parents: 33832
diff changeset
    26
  val type_has_topsort: typ -> bool
37500
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
    27
  val cnf_rules_pairs : theory -> (string * thm) list -> cnf_thm list
37496
9ae78e12e126 reintroduce new Sledgehammer polymorphic handling code;
blanchet
parents: 37488
diff changeset
    28
  val saturate_skolem_cache: theory -> theory option
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37496
diff changeset
    29
  val auto_saturate_skolem_cache: bool Unsynchronized.ref
36398
de8522a5cbae make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents: 36394
diff changeset
    30
  val neg_clausify: thm -> thm list
de8522a5cbae make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents: 36394
diff changeset
    31
  val neg_conjecture_clauses:
de8522a5cbae make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents: 36394
diff changeset
    32
    Proof.context -> thm -> int -> thm list list * (string * typ) list
24669
4579eac2c997 proper signature constraint;
wenzelm
parents: 24632
diff changeset
    33
  val setup: theory -> theory
21505
13d4dba99337 prefer Proof.context over Context.generic;
wenzelm
parents: 21470
diff changeset
    34
end;
19196
62ee8c10d796 Added functions to retrieve local and global atpset rules.
mengj
parents: 19175
diff changeset
    35
37574
b8c1f4c46983 renamed "Sledgehammer_Fact_Preprocessor" to "Clausifier";
blanchet
parents: 37572
diff changeset
    36
structure Clausifier : CLAUSIFIER =
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    37
struct
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    38
37500
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
    39
type cnf_thm = thm * ((string * int) * thm)
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
    40
37541
a76ace919f1c never include anything from the Sledgehammer theory in the relevant facts + killed two obsolete facts
blanchet
parents: 37540
diff changeset
    41
val sledgehammer_prefix = "Sledgehammer" ^ Long_Name.separator
a76ace919f1c never include anything from the Sledgehammer theory in the relevant facts + killed two obsolete facts
blanchet
parents: 37540
diff changeset
    42
37171
fc1e20373e6a make sure chained facts appear in Isar proofs generated by Sledgehammer -- otherwise the proof won't work
blanchet
parents: 36945
diff changeset
    43
(* Used to label theorems chained into the goal. *)
37541
a76ace919f1c never include anything from the Sledgehammer theory in the relevant facts + killed two obsolete facts
blanchet
parents: 37540
diff changeset
    44
val chained_prefix = sledgehammer_prefix ^ "chained_"
37171
fc1e20373e6a make sure chained facts appear in Isar proofs generated by Sledgehammer -- otherwise the proof won't work
blanchet
parents: 36945
diff changeset
    45
32955
4a78daeb012b local channels for tracing/debugging;
wenzelm
parents: 32740
diff changeset
    46
val trace = Unsynchronized.ref false;
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
    47
fun trace_msg msg = if !trace then tracing (msg ()) else ();
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
    48
37567
02e4ccd512b6 further reduce dependencies on "sledgehammer_fact_filter.ML"
blanchet
parents: 37543
diff changeset
    49
fun name_thms_pair_from_ref ctxt chained_ths xref =
02e4ccd512b6 further reduce dependencies on "sledgehammer_fact_filter.ML"
blanchet
parents: 37543
diff changeset
    50
  let
02e4ccd512b6 further reduce dependencies on "sledgehammer_fact_filter.ML"
blanchet
parents: 37543
diff changeset
    51
    val ths = ProofContext.get_fact ctxt xref
02e4ccd512b6 further reduce dependencies on "sledgehammer_fact_filter.ML"
blanchet
parents: 37543
diff changeset
    52
    val name = Facts.string_of_ref xref
02e4ccd512b6 further reduce dependencies on "sledgehammer_fact_filter.ML"
blanchet
parents: 37543
diff changeset
    53
               |> forall (member Thm.eq_thm chained_ths) ths
02e4ccd512b6 further reduce dependencies on "sledgehammer_fact_filter.ML"
blanchet
parents: 37543
diff changeset
    54
                  ? prefix chained_prefix
02e4ccd512b6 further reduce dependencies on "sledgehammer_fact_filter.ML"
blanchet
parents: 37543
diff changeset
    55
  in (name, ths) end
02e4ccd512b6 further reduce dependencies on "sledgehammer_fact_filter.ML"
blanchet
parents: 37543
diff changeset
    56
37541
a76ace919f1c never include anything from the Sledgehammer theory in the relevant facts + killed two obsolete facts
blanchet
parents: 37540
diff changeset
    57
val skolem_theory_name = sledgehammer_prefix ^ "Sko"
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
    58
val skolem_prefix = "sko_"
36492
60532b9bcd1c save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
blanchet
parents: 36478
diff changeset
    59
val skolem_infix = "$"
32955
4a78daeb012b local channels for tracing/debugging;
wenzelm
parents: 32740
diff changeset
    60
35568
8fbbfc39508f renamed type_has_empty_sort to type_has_topsort -- {} is the full universal sort;
wenzelm
parents: 33832
diff changeset
    61
val type_has_topsort = Term.exists_subtype
8fbbfc39508f renamed type_has_empty_sort to type_has_topsort -- {} is the full universal sort;
wenzelm
parents: 33832
diff changeset
    62
  (fn TFree (_, []) => true
8fbbfc39508f renamed type_has_empty_sort to type_has_topsort -- {} is the full universal sort;
wenzelm
parents: 33832
diff changeset
    63
    | TVar (_, []) => true
8fbbfc39508f renamed type_has_empty_sort to type_has_topsort -- {} is the full universal sort;
wenzelm
parents: 33832
diff changeset
    64
    | _ => false);
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
    65
28544
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
    66
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    67
(**** Transformation of Elimination Rules into First-Order Formulas****)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    68
29064
70a61d58460e more antiquotations;
wenzelm
parents: 28965
diff changeset
    69
val cfalse = cterm_of @{theory HOL} HOLogic.false_const;
70a61d58460e more antiquotations;
wenzelm
parents: 28965
diff changeset
    70
val ctp_false = cterm_of @{theory HOL} (HOLogic.mk_Trueprop HOLogic.false_const);
20461
d689ad772b2c skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents: 20457
diff changeset
    71
21430
77651b6d9d6c New transformation of eliminatino rules: we simply replace the final conclusion variable by False
paulson
parents: 21290
diff changeset
    72
(*Converts an elim-rule into an equivalent theorem that does not have the
77651b6d9d6c New transformation of eliminatino rules: we simply replace the final conclusion variable by False
paulson
parents: 21290
diff changeset
    73
  predicate variable.  Leaves other theorems unchanged.  We simply instantiate the
77651b6d9d6c New transformation of eliminatino rules: we simply replace the final conclusion variable by False
paulson
parents: 21290
diff changeset
    74
  conclusion variable to False.*)
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
    75
fun transform_elim th =
21430
77651b6d9d6c New transformation of eliminatino rules: we simply replace the final conclusion variable by False
paulson
parents: 21290
diff changeset
    76
  case concl_of th of    (*conclusion variable*)
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35870
diff changeset
    77
       @{const Trueprop} $ (v as Var (_, @{typ bool})) =>
29064
70a61d58460e more antiquotations;
wenzelm
parents: 28965
diff changeset
    78
           Thm.instantiate ([], [(cterm_of @{theory HOL} v, cfalse)]) th
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35870
diff changeset
    79
    | v as Var(_, @{typ prop}) =>
29064
70a61d58460e more antiquotations;
wenzelm
parents: 28965
diff changeset
    80
           Thm.instantiate ([], [(cterm_of @{theory HOL} v, ctp_false)]) th
21430
77651b6d9d6c New transformation of eliminatino rules: we simply replace the final conclusion variable by False
paulson
parents: 21290
diff changeset
    81
    | _ => th;
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    82
24742
73b8b42a36b6 removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents: 24669
diff changeset
    83
(*To enforce single-threading*)
73b8b42a36b6 removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents: 24669
diff changeset
    84
exception Clausify_failure of theory;
20461
d689ad772b2c skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents: 20457
diff changeset
    85
28544
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
    86
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
    87
(**** SKOLEMIZATION BY INFERENCE (lcp) ****)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
    88
36492
60532b9bcd1c save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
blanchet
parents: 36478
diff changeset
    89
(*Keep the full complexity of the original name*)
60532b9bcd1c save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
blanchet
parents: 36478
diff changeset
    90
fun flatten_name s = space_implode "_X" (Long_Name.explode s);
60532b9bcd1c save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
blanchet
parents: 36478
diff changeset
    91
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
    92
fun skolem_name thm_name j var_name =
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
    93
  skolem_prefix ^ thm_name ^ "_" ^ Int.toString j ^
36492
60532b9bcd1c save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
blanchet
parents: 36478
diff changeset
    94
  skolem_infix ^ (if var_name = "" then "g" else flatten_name var_name)
60532b9bcd1c save the name of Skolemized variables in Sledgehammer for use in the proof reconstruction code
blanchet
parents: 36478
diff changeset
    95
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
    96
(* Hack: Could return false positives (e.g., a user happens to declare a
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
    97
   constant called "SomeTheory.sko_means_shoe_in_$wedish". *)
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
    98
val is_skolem_const_name =
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
    99
  Long_Name.base_name
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   100
  #> String.isPrefix skolem_prefix andf String.isSubstring skolem_infix
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   101
37572
a899f9506f39 more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents: 37571
diff changeset
   102
(* The number of type arguments of a constant, zero if it's monomorphic. For
a899f9506f39 more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents: 37571
diff changeset
   103
   (instances of) Skolem pseudoconstants, this information is encoded in the
a899f9506f39 more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents: 37571
diff changeset
   104
   constant name. *)
a899f9506f39 more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents: 37571
diff changeset
   105
fun num_type_args thy s =
a899f9506f39 more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents: 37571
diff changeset
   106
  if String.isPrefix skolem_theory_name s then
a899f9506f39 more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents: 37571
diff changeset
   107
    s |> unprefix skolem_theory_name
a899f9506f39 more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents: 37571
diff changeset
   108
      |> space_explode skolem_infix |> hd
a899f9506f39 more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents: 37571
diff changeset
   109
      |> space_explode "_" |> List.last |> Int.fromString |> the
a899f9506f39 more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents: 37571
diff changeset
   110
  else
a899f9506f39 more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents: 37571
diff changeset
   111
    (s, Sign.the_const_type thy s) |> Sign.const_typargs thy |> length
a899f9506f39 more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet
parents: 37571
diff changeset
   112
24742
73b8b42a36b6 removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents: 24669
diff changeset
   113
fun rhs_extra_types lhsT rhs =
73b8b42a36b6 removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents: 24669
diff changeset
   114
  let val lhs_vars = Term.add_tfreesT lhsT []
73b8b42a36b6 removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents: 24669
diff changeset
   115
      fun add_new_TFrees (TFree v) =
24821
cc55041ca8ba skolem_cache: ignore internal theorems -- major speedup;
wenzelm
parents: 24785
diff changeset
   116
            if member (op =) lhs_vars v then I else insert (op =) (TFree v)
cc55041ca8ba skolem_cache: ignore internal theorems -- major speedup;
wenzelm
parents: 24785
diff changeset
   117
        | add_new_TFrees _ = I
24742
73b8b42a36b6 removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents: 24669
diff changeset
   118
      val rhs_consts = fold_aterms (fn Const c => insert (op =) c | _ => I) rhs []
73b8b42a36b6 removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents: 24669
diff changeset
   119
  in fold (#2 #> Term.fold_atyps add_new_TFrees) rhs_consts [] end;
73b8b42a36b6 removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents: 24669
diff changeset
   120
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   121
fun skolem_type_and_args bound_T body =
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   122
  let
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   123
    val args1 = OldTerm.term_frees body
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   124
    val Ts1 = map type_of args1
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   125
    val Ts2 = rhs_extra_types (Ts1 ---> bound_T) body
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   126
    val args2 = map (fn T => Free (gensym "vsk", T)) Ts2
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   127
  in (Ts2 ---> Ts1 ---> bound_T, args2 @ args1) end
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   128
37348
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   129
(* Traverse a theorem, declaring Skolem function definitions. String "s" is the
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   130
   suggested prefix for the Skolem constants. *)
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   131
fun declare_skolem_funs s th thy =
27174
c2c484480f40 declare_skofuns/skolem: canonical argument order;
wenzelm
parents: 26939
diff changeset
   132
  let
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   133
    val skolem_count = Unsynchronized.ref 0    (* FIXME ??? *)
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   134
    fun dec_sko (Const (@{const_name Ex}, _) $ (body as Abs (s', T, p)))
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   135
                (axs, thy) =
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   136
        (*Existential: declare a Skolem function, then insert into body and continue*)
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   137
        let
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   138
          val id = skolem_name s (Unsynchronized.inc skolem_count) s'
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   139
          val (cT, args) = skolem_type_and_args T body
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   140
          val rhs = list_abs_free (map dest_Free args,
37410
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37403
diff changeset
   141
                                   HOLogic.choice_const T $ body)
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   142
                  (*Forms a lambda-abstraction over the formal parameters*)
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   143
          val (c, thy) =
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   144
            Sign.declare_const ((Binding.conceal (Binding.name id), cT), NoSyn) thy
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   145
          val cdef = id ^ "_def"
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   146
          val ((_, ax), thy) =
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   147
            Thm.add_def true false (Binding.name cdef, Logic.mk_equals (c, rhs)) thy
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   148
          val ax' = Drule.export_without_context ax
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   149
        in dec_sko (subst_bound (list_comb (c, args), p)) (ax' :: axs, thy) end
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35870
diff changeset
   150
      | dec_sko (Const (@{const_name All}, _) $ (Abs (a, T, p))) thx =
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   151
        (*Universal quant: insert a free variable into body and continue*)
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   152
        let val fname = Name.variant (OldTerm.add_term_names (p, [])) a
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   153
        in dec_sko (subst_bound (Free (fname, T), p)) thx end
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35870
diff changeset
   154
      | dec_sko (@{const "op &"} $ p $ q) thx = dec_sko q (dec_sko p thx)
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35870
diff changeset
   155
      | dec_sko (@{const "op |"} $ p $ q) thx = dec_sko q (dec_sko p thx)
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35870
diff changeset
   156
      | dec_sko (@{const Trueprop} $ p) thx = dec_sko p thx
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37496
diff changeset
   157
      | dec_sko _ thx = thx
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   158
  in dec_sko (prop_of th) ([], thy) end
18141
89e2e8bed08f Skolemization by inference, but not quite finished
paulson
parents: 18009
diff changeset
   159
37410
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37403
diff changeset
   160
fun mk_skolem_id t =
37436
2d76997730a6 found missing beta-eta-contraction
blanchet
parents: 37419
diff changeset
   161
  let val T = fastype_of t in
37496
9ae78e12e126 reintroduce new Sledgehammer polymorphic handling code;
blanchet
parents: 37488
diff changeset
   162
    Const (@{const_name skolem_id}, T --> T) $ t
37436
2d76997730a6 found missing beta-eta-contraction
blanchet
parents: 37419
diff changeset
   163
  end
37410
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37403
diff changeset
   164
37512
ff72d3ddc898 this looks like the most appropriate place to do the beta-eta-contraction
blanchet
parents: 37511
diff changeset
   165
fun quasi_beta_eta_contract (Abs (s, T, t')) =
ff72d3ddc898 this looks like the most appropriate place to do the beta-eta-contraction
blanchet
parents: 37511
diff changeset
   166
    Abs (s, T, quasi_beta_eta_contract t')
ff72d3ddc898 this looks like the most appropriate place to do the beta-eta-contraction
blanchet
parents: 37511
diff changeset
   167
  | quasi_beta_eta_contract t = Envir.beta_eta_contract t
ff72d3ddc898 this looks like the most appropriate place to do the beta-eta-contraction
blanchet
parents: 37511
diff changeset
   168
18141
89e2e8bed08f Skolemization by inference, but not quite finished
paulson
parents: 18009
diff changeset
   169
(*Traverse a theorem, accumulating Skolem function definitions.*)
37518
efb0923cc098 use Skolem inlining also in the nonpolymorphic case, as a step toward simplifying the source code
blanchet
parents: 37512
diff changeset
   170
fun assume_skolem_funs s th =
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   171
  let
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   172
    val skolem_count = Unsynchronized.ref 0   (* FIXME ??? *)
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   173
    fun dec_sko (Const (@{const_name Ex}, _) $ (body as Abs (s', T, p))) defs =
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   174
        (*Existential: declare a Skolem function, then insert into body and continue*)
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   175
        let
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   176
          val skos = map (#1 o Logic.dest_equals) defs  (*existing sko fns*)
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   177
          val args = subtract (op =) skos (OldTerm.term_frees body) (*the formal parameters*)
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   178
          val Ts = map type_of args
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   179
          val cT = Ts ---> T (* FIXME: use "skolem_type_and_args" *)
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   180
          val id = skolem_name s (Unsynchronized.inc skolem_count) s'
37518
efb0923cc098 use Skolem inlining also in the nonpolymorphic case, as a step toward simplifying the source code
blanchet
parents: 37512
diff changeset
   181
          val c = Free (id, cT) (* FIXME: needed? ### *)
37500
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
   182
          (* Forms a lambda-abstraction over the formal parameters *)
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
   183
          val rhs =
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
   184
            list_abs_free (map dest_Free args,
37512
ff72d3ddc898 this looks like the most appropriate place to do the beta-eta-contraction
blanchet
parents: 37511
diff changeset
   185
                           HOLogic.choice_const T
ff72d3ddc898 this looks like the most appropriate place to do the beta-eta-contraction
blanchet
parents: 37511
diff changeset
   186
                           $ quasi_beta_eta_contract body)
37518
efb0923cc098 use Skolem inlining also in the nonpolymorphic case, as a step toward simplifying the source code
blanchet
parents: 37512
diff changeset
   187
            |> mk_skolem_id
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   188
          val def = Logic.mk_equals (c, rhs)
37518
efb0923cc098 use Skolem inlining also in the nonpolymorphic case, as a step toward simplifying the source code
blanchet
parents: 37512
diff changeset
   189
          val comb = list_comb (rhs, args)
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   190
        in dec_sko (subst_bound (comb, p)) (def :: defs) end
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   191
      | dec_sko (Const (@{const_name All},_) $ Abs (a, T, p)) defs =
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   192
        (*Universal quant: insert a free variable into body and continue*)
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   193
        let val fname = Name.variant (OldTerm.add_term_names (p,[])) a
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   194
        in dec_sko (subst_bound (Free(fname,T), p)) defs end
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   195
      | dec_sko (@{const "op &"} $ p $ q) defs = dec_sko q (dec_sko p defs)
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   196
      | dec_sko (@{const "op |"} $ p $ q) defs = dec_sko q (dec_sko p defs)
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   197
      | dec_sko (@{const Trueprop} $ p) defs = dec_sko p defs
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37496
diff changeset
   198
      | dec_sko _ defs = defs
20419
df257a9cf0e9 abstraction of lambda-expressions
paulson
parents: 20373
diff changeset
   199
  in  dec_sko (prop_of th) []  end;
df257a9cf0e9 abstraction of lambda-expressions
paulson
parents: 20373
diff changeset
   200
df257a9cf0e9 abstraction of lambda-expressions
paulson
parents: 20373
diff changeset
   201
24827
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   202
(**** REPLACING ABSTRACTIONS BY COMBINATORS ****)
20419
df257a9cf0e9 abstraction of lambda-expressions
paulson
parents: 20373
diff changeset
   203
df257a9cf0e9 abstraction of lambda-expressions
paulson
parents: 20373
diff changeset
   204
(*Returns the vars of a theorem*)
df257a9cf0e9 abstraction of lambda-expressions
paulson
parents: 20373
diff changeset
   205
fun vars_of_thm th =
22691
290454649b8c Thm.fold_terms;
wenzelm
parents: 22644
diff changeset
   206
  map (Thm.cterm_of (theory_of_thm th) o Var) (Thm.fold_terms Term.add_vars th []);
20419
df257a9cf0e9 abstraction of lambda-expressions
paulson
parents: 20373
diff changeset
   207
37540
48273d1ea331 better eta-expansion in Sledgehammer's clausification;
blanchet
parents: 37518
diff changeset
   208
val fun_cong_all = @{thm expand_fun_eq [THEN iffD1]}
20419
df257a9cf0e9 abstraction of lambda-expressions
paulson
parents: 20373
diff changeset
   209
37540
48273d1ea331 better eta-expansion in Sledgehammer's clausification;
blanchet
parents: 37518
diff changeset
   210
(* Removes the lambdas from an equation of the form t = (%x. u). *)
48273d1ea331 better eta-expansion in Sledgehammer's clausification;
blanchet
parents: 37518
diff changeset
   211
fun extensionalize th =
48273d1ea331 better eta-expansion in Sledgehammer's clausification;
blanchet
parents: 37518
diff changeset
   212
  case prop_of th of
48273d1ea331 better eta-expansion in Sledgehammer's clausification;
blanchet
parents: 37518
diff changeset
   213
    _ $ (Const (@{const_name "op ="}, Type (_, [Type (@{type_name fun}, _), _]))
48273d1ea331 better eta-expansion in Sledgehammer's clausification;
blanchet
parents: 37518
diff changeset
   214
         $ _ $ Abs (s, _, _)) => extensionalize (th RS fun_cong_all)
48273d1ea331 better eta-expansion in Sledgehammer's clausification;
blanchet
parents: 37518
diff changeset
   215
  | _ => th
20419
df257a9cf0e9 abstraction of lambda-expressions
paulson
parents: 20373
diff changeset
   216
37416
d5ac8280497e no point in introducing combinators for inlined Skolem functions
blanchet
parents: 37410
diff changeset
   217
fun is_quasi_lambda_free (Const (@{const_name skolem_id}, _) $ _) = true
d5ac8280497e no point in introducing combinators for inlined Skolem functions
blanchet
parents: 37410
diff changeset
   218
  | is_quasi_lambda_free (t1 $ t2) =
d5ac8280497e no point in introducing combinators for inlined Skolem functions
blanchet
parents: 37410
diff changeset
   219
    is_quasi_lambda_free t1 andalso is_quasi_lambda_free t2
d5ac8280497e no point in introducing combinators for inlined Skolem functions
blanchet
parents: 37410
diff changeset
   220
  | is_quasi_lambda_free (Abs _) = false
d5ac8280497e no point in introducing combinators for inlined Skolem functions
blanchet
parents: 37410
diff changeset
   221
  | is_quasi_lambda_free _ = true
20461
d689ad772b2c skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents: 20457
diff changeset
   222
32010
cb1a1c94b4cd more antiquotations;
wenzelm
parents: 31794
diff changeset
   223
val [f_B,g_B] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_B}));
cb1a1c94b4cd more antiquotations;
wenzelm
parents: 31794
diff changeset
   224
val [g_C,f_C] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_C}));
cb1a1c94b4cd more antiquotations;
wenzelm
parents: 31794
diff changeset
   225
val [f_S,g_S] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_S}));
20863
4ee61dbf192d improvements to abstraction, ensuring more re-use of abstraction functions
paulson
parents: 20789
diff changeset
   226
24827
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   227
(*FIXME: requires more use of cterm constructors*)
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   228
fun abstract ct =
28544
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   229
  let
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   230
      val thy = theory_of_cterm ct
25256
fe467fdf129a Catch exceptions arising during the abstraction operation.
paulson
parents: 25243
diff changeset
   231
      val Abs(x,_,body) = term_of ct
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35870
diff changeset
   232
      val Type(@{type_name fun}, [xT,bodyT]) = typ_of (ctyp_of_term ct)
24827
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   233
      val cxT = ctyp_of thy xT and cbodyT = ctyp_of thy bodyT
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   234
      fun makeK() = instantiate' [SOME cxT, SOME cbodyT] [SOME (cterm_of thy body)] @{thm abs_K}
24827
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   235
  in
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   236
      case body of
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   237
          Const _ => makeK()
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   238
        | Free _ => makeK()
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   239
        | Var _ => makeK()  (*though Var isn't expected*)
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   240
        | Bound 0 => instantiate' [SOME cxT] [] @{thm abs_I} (*identity: I*)
24827
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   241
        | rator$rand =>
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   242
            if loose_bvar1 (rator,0) then (*C or S*)
27179
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   243
               if loose_bvar1 (rand,0) then (*S*)
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   244
                 let val crator = cterm_of thy (Abs(x,xT,rator))
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   245
                     val crand = cterm_of thy (Abs(x,xT,rand))
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   246
                     val abs_S' = cterm_instantiate [(f_S,crator),(g_S,crand)] @{thm abs_S}
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   247
                     val (_,rhs) = Thm.dest_equals (cprop_of abs_S')
27179
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   248
                 in
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   249
                   Thm.transitive abs_S' (Conv.binop_conv abstract rhs)
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   250
                 end
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   251
               else (*C*)
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   252
                 let val crator = cterm_of thy (Abs(x,xT,rator))
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   253
                     val abs_C' = cterm_instantiate [(f_C,crator),(g_C,cterm_of thy rand)] @{thm abs_C}
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   254
                     val (_,rhs) = Thm.dest_equals (cprop_of abs_C')
27179
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   255
                 in
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   256
                   Thm.transitive abs_C' (Conv.fun_conv (Conv.arg_conv abstract) rhs)
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   257
                 end
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   258
            else if loose_bvar1 (rand,0) then (*B or eta*)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36603
diff changeset
   259
               if rand = Bound 0 then Thm.eta_conversion ct
27179
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   260
               else (*B*)
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   261
                 let val crand = cterm_of thy (Abs(x,xT,rand))
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   262
                     val crator = cterm_of thy rator
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   263
                     val abs_B' = cterm_instantiate [(f_B,crator),(g_B,crand)] @{thm abs_B}
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   264
                     val (_,rhs) = Thm.dest_equals (cprop_of abs_B')
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   265
                 in Thm.transitive abs_B' (Conv.arg_conv abstract rhs) end
27179
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   266
            else makeK()
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   267
        | _ => raise Fail "abstract: Bad term"
24827
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   268
  end;
20863
4ee61dbf192d improvements to abstraction, ensuring more re-use of abstraction functions
paulson
parents: 20789
diff changeset
   269
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   270
(* Traverse a theorem, remplacing lambda-abstractions with combinators. *)
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   271
fun do_introduce_combinators ct =
37416
d5ac8280497e no point in introducing combinators for inlined Skolem functions
blanchet
parents: 37410
diff changeset
   272
  if is_quasi_lambda_free (term_of ct) then
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   273
    Thm.reflexive ct
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   274
  else case term_of ct of
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   275
    Abs _ =>
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   276
    let
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   277
      val (cv, cta) = Thm.dest_abs NONE ct
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   278
      val (v, _) = dest_Free (term_of cv)
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   279
      val u_th = do_introduce_combinators cta
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   280
      val cu = Thm.rhs_of u_th
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   281
      val comb_eq = abstract (Thm.cabs cv cu)
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   282
    in Thm.transitive (Thm.abstract_rule v cv u_th) comb_eq end
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   283
  | _ $ _ =>
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   284
    let val (ct1, ct2) = Thm.dest_comb ct in
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   285
        Thm.combination (do_introduce_combinators ct1)
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   286
                        (do_introduce_combinators ct2)
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   287
    end
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   288
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   289
fun introduce_combinators th =
37416
d5ac8280497e no point in introducing combinators for inlined Skolem functions
blanchet
parents: 37410
diff changeset
   290
  if is_quasi_lambda_free (prop_of th) then
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   291
    th
24827
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   292
  else
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   293
    let
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   294
      val th = Drule.eta_contraction_rule th
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   295
      val eqth = do_introduce_combinators (cprop_of th)
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   296
    in Thm.equal_elim eqth th end
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   297
    handle THM (msg, _, _) =>
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   298
           (warning ("Error in the combinator translation of " ^
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   299
                     Display.string_of_thm_without_context th ^
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   300
                     "\nException message: " ^ msg ^ ".");
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   301
            (* A type variable of sort "{}" will make abstraction fail. *)
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   302
            TrueI)
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   303
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   304
(*cterms are used throughout for efficiency*)
29064
70a61d58460e more antiquotations;
wenzelm
parents: 28965
diff changeset
   305
val cTrueprop = Thm.cterm_of @{theory HOL} HOLogic.Trueprop;
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   306
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   307
(*cterm version of mk_cTrueprop*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   308
fun c_mkTrueprop A = Thm.capply cTrueprop A;
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   309
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   310
(*Given an abstraction over n variables, replace the bound variables by free
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   311
  ones. Return the body, along with the list of free variables.*)
20461
d689ad772b2c skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents: 20457
diff changeset
   312
fun c_variant_abs_multi (ct0, vars) =
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   313
      let val (cv,ct) = Thm.dest_abs NONE ct0
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   314
      in  c_variant_abs_multi (ct, cv::vars)  end
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   315
      handle CTERM _ => (ct0, rev vars);
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   316
20461
d689ad772b2c skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents: 20457
diff changeset
   317
(*Given the definition of a Skolem function, return a theorem to replace
d689ad772b2c skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents: 20457
diff changeset
   318
  an existential formula by a use of that function.
18141
89e2e8bed08f Skolemization by inference, but not quite finished
paulson
parents: 18009
diff changeset
   319
   Example: "EX x. x : A & x ~: B ==> sko A B : A & sko A B ~: B"  [.] *)
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   320
fun skolem_theorem_of_def inline def =
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   321
  let
37567
02e4ccd512b6 further reduce dependencies on "sledgehammer_fact_filter.ML"
blanchet
parents: 37543
diff changeset
   322
      val (c, rhs) = def |> Drule.legacy_freeze_thaw |> #1 |> cprop_of
02e4ccd512b6 further reduce dependencies on "sledgehammer_fact_filter.ML"
blanchet
parents: 37543
diff changeset
   323
                         |> Thm.dest_equals
37410
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37403
diff changeset
   324
      val rhs' = rhs |> inline ? (snd o Thm.dest_comb)
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37403
diff changeset
   325
      val (ch, frees) = c_variant_abs_multi (rhs', [])
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37403
diff changeset
   326
      val (chilbert, cabs) = ch |> Thm.dest_comb
26627
dac6d56b7c8d rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26618
diff changeset
   327
      val thy = Thm.theory_of_cterm chilbert
dac6d56b7c8d rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26618
diff changeset
   328
      val t = Thm.term_of chilbert
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   329
      val T =
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   330
        case t of
37410
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37403
diff changeset
   331
          Const (@{const_name Eps}, Type (@{type_name fun}, [_, T])) => T
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37403
diff changeset
   332
        | _ => raise TERM ("skolem_theorem_of_def: expected \"Eps\"", [t])
22596
d0d2af4db18f rep_thm/cterm/ctyp: removed obsolete sign field;
wenzelm
parents: 22516
diff changeset
   333
      val cex = Thm.cterm_of thy (HOLogic.exists_const T)
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   334
      val ex_tm = c_mkTrueprop (Thm.capply cex cabs)
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   335
      and conc =
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   336
        Drule.list_comb (if inline then rhs else c, frees)
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   337
        |> Drule.beta_conv cabs |> c_mkTrueprop
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   338
      fun tacf [prem] =
37410
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37403
diff changeset
   339
        (if inline then rewrite_goals_tac @{thms skolem_id_def_raw}
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37403
diff changeset
   340
         else rewrite_goals_tac [def])
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37403
diff changeset
   341
        THEN rtac ((prem |> inline ? rewrite_rule @{thms skolem_id_def_raw})
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37403
diff changeset
   342
                   RS @{thm someI_ex}) 1
23352
356edb5eb1c4 renamed Goal.prove_raw to Goal.prove_internal;
wenzelm
parents: 22902
diff changeset
   343
  in  Goal.prove_internal [ex_tm] conc tacf
18141
89e2e8bed08f Skolemization by inference, but not quite finished
paulson
parents: 18009
diff changeset
   344
       |> forall_intr_list frees
26653
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26627
diff changeset
   345
       |> Thm.forall_elim_vars 0  (*Introduce Vars, but don't discharge defs.*)
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35828
diff changeset
   346
       |> Thm.varifyT_global
18141
89e2e8bed08f Skolemization by inference, but not quite finished
paulson
parents: 18009
diff changeset
   347
  end;
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   348
24742
73b8b42a36b6 removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents: 24669
diff changeset
   349
20863
4ee61dbf192d improvements to abstraction, ensuring more re-use of abstraction functions
paulson
parents: 20789
diff changeset
   350
(*Converts an Isabelle theorem (intro, elim or simp format, even higher-order) into NNF.*)
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24854
diff changeset
   351
fun to_nnf th ctxt0 =
27179
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   352
  let val th1 = th |> transform_elim |> zero_var_indexes
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32257
diff changeset
   353
      val ((_, [th2]), ctxt) = Variable.import true [th1] ctxt0
37540
48273d1ea331 better eta-expansion in Sledgehammer's clausification;
blanchet
parents: 37518
diff changeset
   354
      val th3 = th2 |> Conv.fconv_rule Object_Logic.atomize
48273d1ea331 better eta-expansion in Sledgehammer's clausification;
blanchet
parents: 37518
diff changeset
   355
                    |> extensionalize
48273d1ea331 better eta-expansion in Sledgehammer's clausification;
blanchet
parents: 37518
diff changeset
   356
                    |> Meson.make_nnf ctxt
24937
340523598914 context-based treatment of generalization; also handling TFrees in axiom clauses
paulson
parents: 24854
diff changeset
   357
  in  (th3, ctxt)  end;
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   358
18141
89e2e8bed08f Skolemization by inference, but not quite finished
paulson
parents: 18009
diff changeset
   359
(*Generate Skolem functions for a theorem supplied in nnf*)
37518
efb0923cc098 use Skolem inlining also in the nonpolymorphic case, as a step toward simplifying the source code
blanchet
parents: 37512
diff changeset
   360
fun skolem_theorems_of_assume s th =
efb0923cc098 use Skolem inlining also in the nonpolymorphic case, as a step toward simplifying the source code
blanchet
parents: 37512
diff changeset
   361
  map (skolem_theorem_of_def true o Thm.assume o cterm_of (theory_of_thm th))
efb0923cc098 use Skolem inlining also in the nonpolymorphic case, as a step toward simplifying the source code
blanchet
parents: 37512
diff changeset
   362
      (assume_skolem_funs s th)
18141
89e2e8bed08f Skolemization by inference, but not quite finished
paulson
parents: 18009
diff changeset
   363
25007
cd497f6fe8d1 trying to make it run faster
paulson
parents: 24959
diff changeset
   364
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   365
(*** Blacklisting (more in "Sledgehammer_Fact_Filter") ***)
25007
cd497f6fe8d1 trying to make it run faster
paulson
parents: 24959
diff changeset
   366
37348
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   367
val max_lambda_nesting = 3
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   368
37348
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   369
fun term_has_too_many_lambdas max (t1 $ t2) =
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   370
    exists (term_has_too_many_lambdas max) [t1, t2]
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   371
  | term_has_too_many_lambdas max (Abs (_, _, t)) =
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   372
    max = 0 orelse term_has_too_many_lambdas (max - 1) t
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   373
  | term_has_too_many_lambdas _ _ = false
25007
cd497f6fe8d1 trying to make it run faster
paulson
parents: 24959
diff changeset
   374
37348
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   375
fun is_formula_type T = (T = HOLogic.boolT orelse T = propT)
25007
cd497f6fe8d1 trying to make it run faster
paulson
parents: 24959
diff changeset
   376
37348
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   377
(* Don't count nested lambdas at the level of formulas, since they are
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   378
   quantifiers. *)
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   379
fun formula_has_too_many_lambdas Ts (Abs (_, T, t)) =
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   380
    formula_has_too_many_lambdas (T :: Ts) t
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   381
  | formula_has_too_many_lambdas Ts t =
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   382
    if is_formula_type (fastype_of1 (Ts, t)) then
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   383
      exists (formula_has_too_many_lambdas Ts) (#2 (strip_comb t))
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   384
    else
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   385
      term_has_too_many_lambdas max_lambda_nesting t
25007
cd497f6fe8d1 trying to make it run faster
paulson
parents: 24959
diff changeset
   386
37348
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   387
(* The max apply depth of any "metis" call in "Metis_Examples" (on 31-10-2007)
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   388
   was 11. *)
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   389
val max_apply_depth = 15
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   390
37348
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   391
fun apply_depth (f $ t) = Int.max (apply_depth f, apply_depth t + 1)
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   392
  | apply_depth (Abs (_, _, t)) = apply_depth t
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   393
  | apply_depth _ = 0
25256
fe467fdf129a Catch exceptions arising during the abstraction operation.
paulson
parents: 25243
diff changeset
   394
37348
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   395
fun is_formula_too_complex t =
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   396
  apply_depth t > max_apply_depth orelse Meson.too_many_clauses NONE t orelse
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   397
  formula_has_too_many_lambdas [] t
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   398
25243
78f8aaa27493 bugfixes concerning strange theorems
paulson
parents: 25007
diff changeset
   399
fun is_strange_thm th =
78f8aaa27493 bugfixes concerning strange theorems
paulson
parents: 25007
diff changeset
   400
  case head_of (concl_of th) of
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35870
diff changeset
   401
      Const (a, _) => (a <> @{const_name Trueprop} andalso
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35870
diff changeset
   402
                       a <> @{const_name "=="})
25243
78f8aaa27493 bugfixes concerning strange theorems
paulson
parents: 25007
diff changeset
   403
    | _ => false;
78f8aaa27493 bugfixes concerning strange theorems
paulson
parents: 25007
diff changeset
   404
37348
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   405
fun is_theorem_bad_for_atps thm =
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   406
  let val t = prop_of thm in
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   407
    is_formula_too_complex t orelse exists_type type_has_topsort t orelse
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   408
    is_strange_thm thm
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   409
  end
25243
78f8aaa27493 bugfixes concerning strange theorems
paulson
parents: 25007
diff changeset
   410
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35870
diff changeset
   411
(* FIXME: put other record thms here, or declare as "no_atp" *)
25007
cd497f6fe8d1 trying to make it run faster
paulson
parents: 24959
diff changeset
   412
val multi_base_blacklist =
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35870
diff changeset
   413
  ["defs", "select_defs", "update_defs", "induct", "inducts", "split", "splits",
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35870
diff changeset
   414
   "split_asm", "cases", "ext_cases"];
25007
cd497f6fe8d1 trying to make it run faster
paulson
parents: 24959
diff changeset
   415
22731
abfdccaed085 trying to make single-step proofs work better, especially if they contain
paulson
parents: 22724
diff changeset
   416
fun fake_name th =
27865
27a8ad9612a3 moved basic thm operations from structure PureThy to Thm (cf. more_thm.ML);
wenzelm
parents: 27809
diff changeset
   417
  if Thm.has_name_hint th then flatten_name (Thm.get_name_hint th)
22731
abfdccaed085 trying to make single-step proofs work better, especially if they contain
paulson
parents: 22724
diff changeset
   418
  else gensym "unknown_thm_";
abfdccaed085 trying to make single-step proofs work better, especially if they contain
paulson
parents: 22724
diff changeset
   419
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   420
(*Skolemize a named theorem, with Skolem functions as additional premises.*)
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   421
fun skolemize_theorem s th =
37345
4402a2bfa204 make Sledgehammer's "add:" and "del:" syntax work better in the presence of aliases;
blanchet
parents: 37332
diff changeset
   422
  if member (op =) multi_base_blacklist (Long_Name.base_name s) orelse
37348
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   423
     is_theorem_bad_for_atps th then
37345
4402a2bfa204 make Sledgehammer's "add:" and "del:" syntax work better in the presence of aliases;
blanchet
parents: 37332
diff changeset
   424
    []
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   425
  else
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   426
    let
36603
d5d6111761a6 renamed Variable.thm_context to Variable.global_thm_context to emphasize that this is not the real thing;
wenzelm
parents: 36492
diff changeset
   427
      val ctxt0 = Variable.global_thm_context th
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   428
      val (nnfth, ctxt) = to_nnf th ctxt0
37518
efb0923cc098 use Skolem inlining also in the nonpolymorphic case, as a step toward simplifying the source code
blanchet
parents: 37512
diff changeset
   429
      val defs = skolem_theorems_of_assume s nnfth
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   430
      val (cnfs, ctxt) = Meson.make_cnf defs nnfth ctxt
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   431
    in
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   432
      cnfs |> map introduce_combinators
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   433
           |> Variable.export ctxt ctxt0
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   434
           |> Meson.finish_cnf
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   435
    end
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   436
    handle THM _ => []
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   437
24742
73b8b42a36b6 removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents: 24669
diff changeset
   438
(*The cache prevents repeated clausification of a theorem, and also repeated declaration of
73b8b42a36b6 removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents: 24669
diff changeset
   439
  Skolem functions.*)
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33339
diff changeset
   440
structure ThmCache = Theory_Data
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22731
diff changeset
   441
(
28544
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   442
  type T = thm list Thmtab.table * unit Symtab.table;
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   443
  val empty = (Thmtab.empty, Symtab.empty);
26618
wenzelm
parents: 26562
diff changeset
   444
  val extend = I;
33522
737589bb9bb8 adapted Theory_Data;
wenzelm
parents: 33339
diff changeset
   445
  fun merge ((cache1, seen1), (cache2, seen2)) : T =
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   446
    (Thmtab.merge (K true) (cache1, cache2), Symtab.merge (K true) (seen1, seen2));
22846
fb79144af9a3 simplified DataFun interfaces;
wenzelm
parents: 22731
diff changeset
   447
);
22516
c986140356b8 Clause cache is now in theory data.
paulson
parents: 22471
diff changeset
   448
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   449
val lookup_cache = Thmtab.lookup o #1 o ThmCache.get;
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   450
val already_seen = Symtab.defined o #2 o ThmCache.get;
20461
d689ad772b2c skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents: 20457
diff changeset
   451
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   452
val update_cache = ThmCache.map o apfst o Thmtab.update;
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   453
fun mark_seen name = ThmCache.map (apsnd (Symtab.update (name, ())));
25007
cd497f6fe8d1 trying to make it run faster
paulson
parents: 24959
diff changeset
   454
36228
df47dc6c0e03 got rid of somewhat pointless "pairname" function in Sledgehammer
blanchet
parents: 36106
diff changeset
   455
(* Convert Isabelle theorems into axiom clauses. *)
27179
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   456
fun cnf_axiom thy th0 =
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   457
  let val th = Thm.transfer thy th0 in
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   458
    case lookup_cache thy th of
37416
d5ac8280497e no point in introducing combinators for inlined Skolem functions
blanchet
parents: 37410
diff changeset
   459
      SOME cls => cls
d5ac8280497e no point in introducing combinators for inlined Skolem functions
blanchet
parents: 37410
diff changeset
   460
    | NONE => map Thm.close_derivation (skolemize_theorem (fake_name th) th)
22516
c986140356b8 Clause cache is now in theory data.
paulson
parents: 22471
diff changeset
   461
  end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   462
18141
89e2e8bed08f Skolemization by inference, but not quite finished
paulson
parents: 18009
diff changeset
   463
22471
7c51f1a799f3 Removal of axiom names from the theorem cache
paulson
parents: 22360
diff changeset
   464
(**** Translate a set of theorems into CNF ****)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   465
21290
33b6bb5d6ab8 Improvement to classrel clauses: now outputs the minimum needed.
paulson
parents: 21254
diff changeset
   466
(*The combination of rev and tail recursion preserves the original order*)
37416
d5ac8280497e no point in introducing combinators for inlined Skolem functions
blanchet
parents: 37410
diff changeset
   467
fun cnf_rules_pairs thy =
d5ac8280497e no point in introducing combinators for inlined Skolem functions
blanchet
parents: 37410
diff changeset
   468
  let
37500
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
   469
    fun do_one _ [] = []
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
   470
      | do_one ((name, k), th) (cls :: clss) =
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
   471
        (cls, ((name, k), th)) :: do_one ((name, k + 1), th) clss
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
   472
    fun do_all pairs [] = pairs
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
   473
      | do_all pairs ((name, th) :: ths) =
37416
d5ac8280497e no point in introducing combinators for inlined Skolem functions
blanchet
parents: 37410
diff changeset
   474
        let
37500
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
   475
          val new_pairs = do_one ((name, 0), th) (cnf_axiom thy th)
37570
de5feddaa95c got rid of needless exception
blanchet
parents: 37567
diff changeset
   476
                          handle THM _ => []
37500
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
   477
        in do_all (new_pairs @ pairs) ths end
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
   478
  in do_all [] o rev end
19353
36b6b15ee670 added another function for CNF.
mengj
parents: 19232
diff changeset
   479
19196
62ee8c10d796 Added functions to retrieve local and global atpset rules.
mengj
parents: 19175
diff changeset
   480
35865
2f8fb5242799 more Sledgehammer refactoring
blanchet
parents: 35828
diff changeset
   481
(**** Convert all facts of the theory into FOL or HOL clauses ****)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   482
28544
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   483
local
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   484
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   485
fun skolem_def (name, th) thy =
36603
d5d6111761a6 renamed Variable.thm_context to Variable.global_thm_context to emphasize that this is not the real thing;
wenzelm
parents: 36492
diff changeset
   486
  let val ctxt0 = Variable.global_thm_context th in
37348
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   487
    case try (to_nnf th) ctxt0 of
28544
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   488
      NONE => (NONE, thy)
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   489
    | SOME (nnfth, ctxt) =>
37348
3ad1bfd2de46 renaming
blanchet
parents: 37347
diff changeset
   490
      let val (defs, thy') = declare_skolem_funs (flatten_name name) nnfth thy
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   491
      in (SOME (th, ctxt0, ctxt, nnfth, defs), thy') end
28544
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   492
  end;
24742
73b8b42a36b6 removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents: 24669
diff changeset
   493
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   494
fun skolem_cnfs (th, ctxt0, ctxt, nnfth, defs) =
28544
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   495
  let
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   496
    val (cnfs, ctxt) =
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   497
      Meson.make_cnf (map (skolem_theorem_of_def false) defs) nnfth ctxt
28544
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   498
    val cnfs' = cnfs
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   499
      |> map introduce_combinators
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   500
      |> Variable.export ctxt ctxt0
28544
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   501
      |> Meson.finish_cnf
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   502
      |> map Thm.close_derivation;
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   503
    in (th, cnfs') end;
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   504
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   505
in
24742
73b8b42a36b6 removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents: 24669
diff changeset
   506
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   507
fun saturate_skolem_cache thy =
28544
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   508
  let
33306
4138ba02b681 replaced slightly odd Thm.is_internal by Facts.is_concealed -- as provided by the name space;
wenzelm
parents: 33222
diff changeset
   509
    val facts = PureThy.facts_of thy;
4138ba02b681 replaced slightly odd Thm.is_internal by Facts.is_concealed -- as provided by the name space;
wenzelm
parents: 33222
diff changeset
   510
    val new_facts = (facts, []) |-> Facts.fold_static (fn (name, ths) =>
4138ba02b681 replaced slightly odd Thm.is_internal by Facts.is_concealed -- as provided by the name space;
wenzelm
parents: 33222
diff changeset
   511
      if Facts.is_concealed facts name orelse already_seen thy name then I
4138ba02b681 replaced slightly odd Thm.is_internal by Facts.is_concealed -- as provided by the name space;
wenzelm
parents: 33222
diff changeset
   512
      else cons (name, ths));
28544
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   513
    val new_thms = (new_facts, []) |-> fold (fn (name, ths) =>
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   514
      if member (op =) multi_base_blacklist (Long_Name.base_name name) then
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   515
        I
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   516
      else
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   517
        fold_index (fn (i, th) =>
37410
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37403
diff changeset
   518
          if is_theorem_bad_for_atps th orelse
2bf7e6136047 adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents: 37403
diff changeset
   519
             is_some (lookup_cache thy th) then
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   520
            I
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   521
          else
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   522
            cons (name ^ "_" ^ string_of_int (i + 1), Thm.transfer thy th)) ths)
28544
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   523
  in
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   524
    if null new_facts then
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   525
      NONE
28544
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   526
    else
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   527
      let
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   528
        val (defs, thy') = thy
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   529
          |> fold (mark_seen o #1) new_facts
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   530
          |> fold_map skolem_def (sort_distinct (Thm.thm_ord o pairself snd) new_thms)
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   531
          |>> map_filter I;
29368
503ce3f8f092 renamed structure ParList to Par_List;
wenzelm
parents: 29270
diff changeset
   532
        val cache_entries = Par_List.map skolem_cnfs defs;
28544
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   533
      in SOME (fold update_cache cache_entries thy') end
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   534
  end;
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   535
28544
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   536
end;
24854
0ebcd575d3c6 filtering out some package theorems
paulson
parents: 24827
diff changeset
   537
37511
26afa11a1fb2 killed legacy "neg_clausify" and "clausify"
blanchet
parents: 37500
diff changeset
   538
(* For emergency use where the Skolem cache causes problems. *)
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37496
diff changeset
   539
val auto_saturate_skolem_cache = Unsynchronized.ref true
20457
85414caac94a refinements to conversion into clause form, esp for the HO case
paulson
parents: 20445
diff changeset
   540
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37496
diff changeset
   541
fun conditionally_saturate_skolem_cache thy =
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37496
diff changeset
   542
  if !auto_saturate_skolem_cache then saturate_skolem_cache thy else NONE
27179
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   543
36398
de8522a5cbae make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents: 36394
diff changeset
   544
21999
0cf192e489e2 improvements to proof reconstruction. Some files loaded in a different order
paulson
parents: 21900
diff changeset
   545
(*** Converting a subgoal into negated conjecture clauses. ***)
0cf192e489e2 improvements to proof reconstruction. Some files loaded in a different order
paulson
parents: 21900
diff changeset
   546
32262
73cd8f74cf2a Meson.first_order_resolve: avoid handle _;
wenzelm
parents: 32257
diff changeset
   547
fun neg_skolemize_tac ctxt =
37332
51d99ba6fc4d don't raise Option.Option if assumptions contain schematic variables
blanchet
parents: 37171
diff changeset
   548
  EVERY' [rtac ccontr, Object_Logic.atomize_prems_tac, Meson.skolemize_tac ctxt]
36398
de8522a5cbae make "neg_clausify" return a list of lists of clauses, so that it's possible to see which clause comes from which theorem
blanchet
parents: 36394
diff changeset
   549
35869
cac366550624 start work on direct proof reconstruction for Sledgehammer
blanchet
parents: 35865
diff changeset
   550
val neg_clausify =
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   551
  single
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   552
  #> Meson.make_clauses_unsorted
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   553
  #> map introduce_combinators
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   554
  #> Meson.finish_cnf
21999
0cf192e489e2 improvements to proof reconstruction. Some files loaded in a different order
paulson
parents: 21900
diff changeset
   555
32257
bad5a99c16d8 neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
wenzelm
parents: 32231
diff changeset
   556
fun neg_conjecture_clauses ctxt st0 n =
bad5a99c16d8 neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
wenzelm
parents: 32231
diff changeset
   557
  let
37332
51d99ba6fc4d don't raise Option.Option if assumptions contain schematic variables
blanchet
parents: 37171
diff changeset
   558
    (* "Option" is thrown if the assumptions contain schematic variables. *)
51d99ba6fc4d don't raise Option.Option if assumptions contain schematic variables
blanchet
parents: 37171
diff changeset
   559
    val st = Seq.hd (neg_skolemize_tac ctxt n st0) handle Option.Option => st0
51d99ba6fc4d don't raise Option.Option if assumptions contain schematic variables
blanchet
parents: 37171
diff changeset
   560
    val ({params, prems, ...}, _) =
51d99ba6fc4d don't raise Option.Option if assumptions contain schematic variables
blanchet
parents: 37171
diff changeset
   561
      Subgoal.focus (Variable.set_body false ctxt) n st
51d99ba6fc4d don't raise Option.Option if assumptions contain schematic variables
blanchet
parents: 37171
diff changeset
   562
  in (map neg_clausify prems, map (dest_Free o term_of o #2) params) end
21999
0cf192e489e2 improvements to proof reconstruction. Some files loaded in a different order
paulson
parents: 21900
diff changeset
   563
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   564
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   565
(** setup **)
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   566
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   567
val setup =
37498
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37496
diff changeset
   568
  perhaps conditionally_saturate_skolem_cache
b426cbdb5a23 removed Sledgehammer's support for the DFG syntax;
blanchet
parents: 37496
diff changeset
   569
  #> Theory.at_end conditionally_saturate_skolem_cache
18510
0a6c24f549c3 the "skolem" attribute and better initialization of the clause database
paulson
parents: 18404
diff changeset
   570
20461
d689ad772b2c skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents: 20457
diff changeset
   571
end;