src/HOL/Tools/Meson/meson_clausify.ML
author blanchet
Tue, 22 May 2012 16:59:27 +0200
changeset 47954 aada9fd08b58
parent 47953 a2c3706c4cb1
child 50705 0e943b33d907
permissions -rw-r--r--
make higher-order goals more first-order via extensionality
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
39941
02fcd9cd1eac move Meson to Plain
blanchet
parents: 39940
diff changeset
     1
(*  Title:      HOL/Tools/Meson/meson_clausify.ML
38027
505657ddb047 standardize "Author" tags
blanchet
parents: 38016
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: 36228
diff changeset
     3
    Author:     Jasmin Blanchette, TU Muenchen
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     4
39941
02fcd9cd1eac move Meson to Plain
blanchet
parents: 39940
diff changeset
     5
Transformation of HOL theorems 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
39890
a1695e2169d0 finished renaming file and module
blanchet
parents: 39889
diff changeset
     8
signature MESON_CLAUSIFY =
21505
13d4dba99337 prefer Proof.context over Context.generic;
wenzelm
parents: 21470
diff changeset
     9
sig
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
    10
  val new_skolem_var_prefix : string
42098
f978caf60bbe more robust handling of variables in new Skolemizer
blanchet
parents: 42072
diff changeset
    11
  val new_nonskolem_var_prefix : string
42099
447fa058ab22 avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents: 42098
diff changeset
    12
  val is_zapped_var_name : string -> bool
45511
9b0f8ca4388e continued implementation of lambda-lifting in Metis
blanchet
parents: 45508
diff changeset
    13
  val is_quasi_lambda_free : term -> bool
38001
a9b47b85ca24 reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents: 38000
diff changeset
    14
  val introduce_combinators_in_cterm : cterm -> thm
38028
22dcaec5fa77 minor refactoring
blanchet
parents: 38027
diff changeset
    15
  val introduce_combinators_in_theorem : thm -> thm
39932
acde1b606b0e reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
blanchet
parents: 39931
diff changeset
    16
  val cluster_of_zapped_var_name : string -> (int * (int * int)) * bool
42336
d63d43e85879 improve definitional CNF on goal by moving "not" past the quantifiers
blanchet
parents: 42335
diff changeset
    17
  val ss_only : thm list -> simpset
39897
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39896
diff changeset
    18
  val cnf_axiom :
45508
b216dc1b3630 started implementing lambda-lifting in Metis
blanchet
parents: 44241
diff changeset
    19
    Proof.context -> bool -> bool -> int -> thm
b216dc1b3630 started implementing lambda-lifting in Metis
blanchet
parents: 44241
diff changeset
    20
    -> (thm * term) option * thm list
21505
13d4dba99337 prefer Proof.context over Context.generic;
wenzelm
parents: 21470
diff changeset
    21
end;
19196
62ee8c10d796 Added functions to retrieve local and global atpset rules.
mengj
parents: 19175
diff changeset
    22
39890
a1695e2169d0 finished renaming file and module
blanchet
parents: 39889
diff changeset
    23
structure Meson_Clausify : MESON_CLAUSIFY =
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    24
struct
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    25
39950
f3c4849868b8 got rid of overkill "meson_choice" attribute;
blanchet
parents: 39949
diff changeset
    26
open Meson
f3c4849868b8 got rid of overkill "meson_choice" attribute;
blanchet
parents: 39949
diff changeset
    27
42072
1492ee6b8085 avoid illegal variable names, which then yield "Not an identifier" errors in "rename_tac"
blanchet
parents: 41225
diff changeset
    28
(* the extra "Meson" helps prevent clashes (FIXME) *)
1492ee6b8085 avoid illegal variable names, which then yield "Not an identifier" errors in "rename_tac"
blanchet
parents: 41225
diff changeset
    29
val new_skolem_var_prefix = "MesonSK"
1492ee6b8085 avoid illegal variable names, which then yield "Not an identifier" errors in "rename_tac"
blanchet
parents: 41225
diff changeset
    30
val new_nonskolem_var_prefix = "MesonV"
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
    31
42099
447fa058ab22 avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents: 42098
diff changeset
    32
fun is_zapped_var_name s =
447fa058ab22 avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents: 42098
diff changeset
    33
  exists (fn prefix => String.isPrefix prefix s)
447fa058ab22 avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents: 42098
diff changeset
    34
         [new_skolem_var_prefix, new_nonskolem_var_prefix]
447fa058ab22 avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents: 42098
diff changeset
    35
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    36
(**** Transformation of Elimination Rules into First-Order Formulas****)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    37
45740
132a3e1c0fe5 more antiquotations;
wenzelm
parents: 45511
diff changeset
    38
val cfalse = cterm_of @{theory HOL} @{term False};
132a3e1c0fe5 more antiquotations;
wenzelm
parents: 45511
diff changeset
    39
val ctp_false = cterm_of @{theory HOL} (HOLogic.mk_Trueprop @{term False});
20461
d689ad772b2c skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents: 20457
diff changeset
    40
38001
a9b47b85ca24 reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents: 38000
diff changeset
    41
(* Converts an elim-rule into an equivalent theorem that does not have the
a9b47b85ca24 reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents: 38000
diff changeset
    42
   predicate variable. Leaves other theorems unchanged. We simply instantiate
42944
9e620869a576 improved Waldmeister support -- even run it by default on unit equational goals
blanchet
parents: 42747
diff changeset
    43
   the conclusion variable to False. (Cf. "transform_elim_prop" in
38652
e063be321438 perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents: 38632
diff changeset
    44
   "Sledgehammer_Util".) *)
38001
a9b47b85ca24 reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents: 38000
diff changeset
    45
fun transform_elim_theorem th =
21430
77651b6d9d6c New transformation of eliminatino rules: we simply replace the final conclusion variable by False
paulson
parents: 21290
diff changeset
    46
  case concl_of th of    (*conclusion variable*)
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35870
diff changeset
    47
       @{const Trueprop} $ (v as Var (_, @{typ bool})) =>
29064
70a61d58460e more antiquotations;
wenzelm
parents: 28965
diff changeset
    48
           Thm.instantiate ([], [(cterm_of @{theory HOL} v, cfalse)]) th
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35870
diff changeset
    49
    | v as Var(_, @{typ prop}) =>
29064
70a61d58460e more antiquotations;
wenzelm
parents: 28965
diff changeset
    50
           Thm.instantiate ([], [(cterm_of @{theory HOL} v, ctp_false)]) th
38001
a9b47b85ca24 reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents: 38000
diff changeset
    51
    | _ => th
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    52
28544
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
    53
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
    54
(**** SKOLEMIZATION BY INFERENCE (lcp) ****)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
    55
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39721
diff changeset
    56
fun mk_old_skolem_term_wrapper t =
37436
2d76997730a6 found missing beta-eta-contraction
blanchet
parents: 37419
diff changeset
    57
  let val T = fastype_of t in
39962
d42ddd7407ca qualify "Meson." and "Metis." constants so that the ML files can be reloaded later (e.g. for testing/development)
blanchet
parents: 39950
diff changeset
    58
    Const (@{const_name Meson.skolem}, T --> T) $ t
37436
2d76997730a6 found missing beta-eta-contraction
blanchet
parents: 37419
diff changeset
    59
  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
    60
39931
97b8051033be renamed internal function
blanchet
parents: 39908
diff changeset
    61
fun beta_eta_in_abs_body (Abs (s, T, t')) = Abs (s, T, beta_eta_in_abs_body t')
97b8051033be renamed internal function
blanchet
parents: 39908
diff changeset
    62
  | beta_eta_in_abs_body t = Envir.beta_eta_contract t
37512
ff72d3ddc898 this looks like the most appropriate place to do the beta-eta-contraction
blanchet
parents: 37511
diff changeset
    63
18141
89e2e8bed08f Skolemization by inference, but not quite finished
paulson
parents: 18009
diff changeset
    64
(*Traverse a theorem, accumulating Skolem function definitions.*)
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39721
diff changeset
    65
fun old_skolem_defs th =
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
    66
  let
39376
blanchet
parents: 39355
diff changeset
    67
    fun dec_sko (Const (@{const_name Ex}, _) $ (body as Abs (_, T, p))) rhss =
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
    68
        (*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
    69
        let
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43964
diff changeset
    70
          val args = Misc_Legacy.term_frees body
37500
7587b6e63454 thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents: 37498
diff changeset
    71
          (* 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
    72
          val rhs =
44241
7943b69f0188 modernized signature of Term.absfree/absdummy;
wenzelm
parents: 44121
diff changeset
    73
            fold_rev (absfree o dest_Free) args
7943b69f0188 modernized signature of Term.absfree/absdummy;
wenzelm
parents: 44121
diff changeset
    74
              (HOLogic.choice_const T $ beta_eta_in_abs_body body)
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39721
diff changeset
    75
            |> mk_old_skolem_term_wrapper
37518
efb0923cc098 use Skolem inlining also in the nonpolymorphic case, as a step toward simplifying the source code
blanchet
parents: 37512
diff changeset
    76
          val comb = list_comb (rhs, args)
37617
f73cd4069f69 always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents: 37616
diff changeset
    77
        in dec_sko (subst_bound (comb, p)) (rhs :: rhss) end
f73cd4069f69 always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents: 37616
diff changeset
    78
      | dec_sko (Const (@{const_name All},_) $ Abs (a, T, p)) rhss =
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
    79
        (*Universal quant: insert a free variable into body and continue*)
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43964
diff changeset
    80
        let val fname = singleton (Name.variant_list (Misc_Legacy.add_term_names (p, []))) a
37617
f73cd4069f69 always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents: 37616
diff changeset
    81
        in dec_sko (subst_bound (Free(fname,T), p)) rhss end
39906
blanchet
parents: 39905
diff changeset
    82
      | dec_sko (@{const conj} $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q
blanchet
parents: 39905
diff changeset
    83
      | dec_sko (@{const disj} $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q
37617
f73cd4069f69 always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents: 37616
diff changeset
    84
      | dec_sko (@{const Trueprop} $ p) rhss = dec_sko p rhss
f73cd4069f69 always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents: 37616
diff changeset
    85
      | dec_sko _ rhss = rhss
20419
df257a9cf0e9 abstraction of lambda-expressions
paulson
parents: 20373
diff changeset
    86
  in  dec_sko (prop_of th) []  end;
df257a9cf0e9 abstraction of lambda-expressions
paulson
parents: 20373
diff changeset
    87
df257a9cf0e9 abstraction of lambda-expressions
paulson
parents: 20373
diff changeset
    88
24827
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
    89
(**** REPLACING ABSTRACTIONS BY COMBINATORS ****)
20419
df257a9cf0e9 abstraction of lambda-expressions
paulson
parents: 20373
diff changeset
    90
39962
d42ddd7407ca qualify "Meson." and "Metis." constants so that the ML files can be reloaded later (e.g. for testing/development)
blanchet
parents: 39950
diff changeset
    91
fun is_quasi_lambda_free (Const (@{const_name Meson.skolem}, _) $ _) = true
37416
d5ac8280497e no point in introducing combinators for inlined Skolem functions
blanchet
parents: 37410
diff changeset
    92
  | is_quasi_lambda_free (t1 $ t2) =
d5ac8280497e no point in introducing combinators for inlined Skolem functions
blanchet
parents: 37410
diff changeset
    93
    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
    94
  | is_quasi_lambda_free (Abs _) = false
d5ac8280497e no point in introducing combinators for inlined Skolem functions
blanchet
parents: 37410
diff changeset
    95
  | is_quasi_lambda_free _ = true
20461
d689ad772b2c skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents: 20457
diff changeset
    96
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43964
diff changeset
    97
val [f_B,g_B] = map (cterm_of @{theory}) (Misc_Legacy.term_vars (prop_of @{thm abs_B}));
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43964
diff changeset
    98
val [g_C,f_C] = map (cterm_of @{theory}) (Misc_Legacy.term_vars (prop_of @{thm abs_C}));
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43964
diff changeset
    99
val [f_S,g_S] = map (cterm_of @{theory}) (Misc_Legacy.term_vars (prop_of @{thm abs_S}));
20863
4ee61dbf192d improvements to abstraction, ensuring more re-use of abstraction functions
paulson
parents: 20789
diff changeset
   100
38282
319c59682c51 move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents: 38280
diff changeset
   101
(* FIXME: Requires more use of cterm constructors. *)
24827
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   102
fun abstract ct =
28544
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   103
  let
26743a1591f5 improved performance of skolem cache, due to parallel map;
wenzelm
parents: 28262
diff changeset
   104
      val thy = theory_of_cterm ct
25256
fe467fdf129a Catch exceptions arising during the abstraction operation.
paulson
parents: 25243
diff changeset
   105
      val Abs(x,_,body) = term_of ct
35963
943e2582dc87 added options to Sledgehammer;
blanchet
parents: 35870
diff changeset
   106
      val Type(@{type_name fun}, [xT,bodyT]) = typ_of (ctyp_of_term ct)
38005
b6555e9c5de4 prevent schematic variable clash in combinator-introduction code, when invoked from Sledgehammer (another consequence of the CNF -> FOF transition)
blanchet
parents: 38001
diff changeset
   107
      val cxT = ctyp_of thy xT
b6555e9c5de4 prevent schematic variable clash in combinator-introduction code, when invoked from Sledgehammer (another consequence of the CNF -> FOF transition)
blanchet
parents: 38001
diff changeset
   108
      val cbodyT = ctyp_of thy bodyT
b6555e9c5de4 prevent schematic variable clash in combinator-introduction code, when invoked from Sledgehammer (another consequence of the CNF -> FOF transition)
blanchet
parents: 38001
diff changeset
   109
      fun makeK () =
b6555e9c5de4 prevent schematic variable clash in combinator-introduction code, when invoked from Sledgehammer (another consequence of the CNF -> FOF transition)
blanchet
parents: 38001
diff changeset
   110
        instantiate' [SOME cxT, SOME cbodyT] [SOME (cterm_of thy body)]
b6555e9c5de4 prevent schematic variable clash in combinator-introduction code, when invoked from Sledgehammer (another consequence of the CNF -> FOF transition)
blanchet
parents: 38001
diff changeset
   111
                     @{thm abs_K}
24827
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   112
  in
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   113
      case body of
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   114
          Const _ => makeK()
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   115
        | Free _ => makeK()
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   116
        | Var _ => makeK()  (*though Var isn't expected*)
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   117
        | Bound 0 => instantiate' [SOME cxT] [] @{thm abs_I} (*identity: I*)
24827
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   118
        | rator$rand =>
42083
e1209fc7ecdc added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
wenzelm
parents: 42072
diff changeset
   119
            if Term.is_dependent rator then (*C or S*)
e1209fc7ecdc added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
wenzelm
parents: 42072
diff changeset
   120
               if Term.is_dependent rand then (*S*)
27179
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   121
                 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
   122
                     val crand = cterm_of thy (Abs(x,xT,rand))
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   123
                     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
   124
                     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
   125
                 in
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   126
                   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
   127
                 end
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   128
               else (*C*)
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   129
                 let val crator = cterm_of thy (Abs(x,xT,rator))
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   130
                     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
   131
                     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
   132
                 in
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   133
                   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
   134
                 end
42083
e1209fc7ecdc added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
wenzelm
parents: 42072
diff changeset
   135
            else if Term.is_dependent rand then (*B or eta*)
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36603
diff changeset
   136
               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
   137
               else (*B*)
8f29fed3dc9a ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents: 27174
diff changeset
   138
                 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
   139
                     val crator = cterm_of thy rator
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   140
                     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
   141
                     val (_,rhs) = Thm.dest_equals (cprop_of abs_B')
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   142
                 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
   143
            else makeK()
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   144
        | _ => raise Fail "abstract: Bad term"
24827
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   145
  end;
20863
4ee61dbf192d improvements to abstraction, ensuring more re-use of abstraction functions
paulson
parents: 20789
diff changeset
   146
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   147
(* Traverse a theorem, remplacing lambda-abstractions with combinators. *)
38001
a9b47b85ca24 reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents: 38000
diff changeset
   148
fun introduce_combinators_in_cterm ct =
37416
d5ac8280497e no point in introducing combinators for inlined Skolem functions
blanchet
parents: 37410
diff changeset
   149
  if is_quasi_lambda_free (term_of ct) then
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   150
    Thm.reflexive ct
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   151
  else case term_of ct of
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   152
    Abs _ =>
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   153
    let
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   154
      val (cv, cta) = Thm.dest_abs NONE ct
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   155
      val (v, _) = dest_Free (term_of cv)
38001
a9b47b85ca24 reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents: 38000
diff changeset
   156
      val u_th = introduce_combinators_in_cterm cta
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   157
      val cu = Thm.rhs_of u_th
46497
89ccf66aa73d renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents: 46071
diff changeset
   158
      val comb_eq = abstract (Thm.lambda cv cu)
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   159
    in Thm.transitive (Thm.abstract_rule v cv u_th) comb_eq end
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   160
  | _ $ _ =>
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   161
    let val (ct1, ct2) = Thm.dest_comb ct in
38001
a9b47b85ca24 reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents: 38000
diff changeset
   162
        Thm.combination (introduce_combinators_in_cterm ct1)
a9b47b85ca24 reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents: 38000
diff changeset
   163
                        (introduce_combinators_in_cterm ct2)
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   164
    end
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   165
38001
a9b47b85ca24 reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents: 38000
diff changeset
   166
fun introduce_combinators_in_theorem th =
37416
d5ac8280497e no point in introducing combinators for inlined Skolem functions
blanchet
parents: 37410
diff changeset
   167
  if is_quasi_lambda_free (prop_of th) then
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   168
    th
24827
646bdc51eb7d combinator translation
paulson
parents: 24821
diff changeset
   169
  else
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   170
    let
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   171
      val th = Drule.eta_contraction_rule th
38001
a9b47b85ca24 reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents: 38000
diff changeset
   172
      val eqth = introduce_combinators_in_cterm (cprop_of th)
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   173
    in Thm.equal_elim eqth th end
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   174
    handle THM (msg, _, _) =>
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   175
           (warning ("Error in the combinator translation of " ^
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   176
                     Display.string_of_thm_without_context th ^
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   177
                     "\nException message: " ^ msg ^ ".");
45511
9b0f8ca4388e continued implementation of lambda-lifting in Metis
blanchet
parents: 45508
diff changeset
   178
            (* A type variable of sort "{}" will make "abstraction" fail. *)
37349
3d7058e24b7a cosmetics
blanchet
parents: 37348
diff changeset
   179
            TrueI)
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   180
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   181
(*cterms are used throughout for efficiency*)
38280
577f138af235 replace recursion with "fold"
blanchet
parents: 38278
diff changeset
   182
val cTrueprop = cterm_of @{theory HOL} HOLogic.Trueprop;
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   183
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   184
(*Given an abstraction over n variables, replace the bound variables by free
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   185
  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
   186
fun c_variant_abs_multi (ct0, vars) =
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   187
      let val (cv,ct) = Thm.dest_abs NONE ct0
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   188
      in  c_variant_abs_multi (ct, cv::vars)  end
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   189
      handle CTERM _ => (ct0, rev vars);
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   190
37617
f73cd4069f69 always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents: 37616
diff changeset
   191
(* Given the definition of a Skolem function, return a theorem to replace
f73cd4069f69 always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents: 37616
diff changeset
   192
   an existential formula by a use of that function.
18141
89e2e8bed08f Skolemization by inference, but not quite finished
paulson
parents: 18009
diff changeset
   193
   Example: "EX x. x : A & x ~: B ==> sko A B : A & sko A B ~: B"  [.] *)
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39721
diff changeset
   194
fun old_skolem_theorem_from_def thy rhs0 =
37399
34f080a12063 proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents: 37349
diff changeset
   195
  let
38280
577f138af235 replace recursion with "fold"
blanchet
parents: 38278
diff changeset
   196
    val rhs = rhs0 |> Type.legacy_freeze_thaw |> #1 |> cterm_of thy
37617
f73cd4069f69 always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents: 37616
diff changeset
   197
    val rhs' = rhs |> Thm.dest_comb |> snd
f73cd4069f69 always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents: 37616
diff changeset
   198
    val (ch, frees) = c_variant_abs_multi (rhs', [])
f73cd4069f69 always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents: 37616
diff changeset
   199
    val (hilbert, cabs) = ch |> Thm.dest_comb |>> term_of
f73cd4069f69 always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents: 37616
diff changeset
   200
    val T =
f73cd4069f69 always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents: 37616
diff changeset
   201
      case hilbert of
39941
02fcd9cd1eac move Meson to Plain
blanchet
parents: 39940
diff changeset
   202
        Const (_, Type (@{type_name fun}, [_, T])) => T
39886
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39721
diff changeset
   203
      | _ => raise TERM ("old_skolem_theorem_from_def: expected \"Eps\"",
8a9f0c97d550 first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents: 39721
diff changeset
   204
                         [hilbert])
38280
577f138af235 replace recursion with "fold"
blanchet
parents: 38278
diff changeset
   205
    val cex = cterm_of thy (HOLogic.exists_const T)
46497
89ccf66aa73d renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents: 46071
diff changeset
   206
    val ex_tm = Thm.apply cTrueprop (Thm.apply cex cabs)
37629
a4f129820562 more elegant cheating
blanchet
parents: 37628
diff changeset
   207
    val conc =
37617
f73cd4069f69 always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents: 37616
diff changeset
   208
      Drule.list_comb (rhs, frees)
46497
89ccf66aa73d renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents: 46071
diff changeset
   209
      |> Drule.beta_conv cabs |> Thm.apply cTrueprop
37617
f73cd4069f69 always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents: 37616
diff changeset
   210
    fun tacf [prem] =
46904
f30e941b4512 prefer abs_def over def_raw;
wenzelm
parents: 46497
diff changeset
   211
      rewrite_goals_tac @{thms skolem_def [abs_def]}
f30e941b4512 prefer abs_def over def_raw;
wenzelm
parents: 46497
diff changeset
   212
      THEN rtac ((prem |> rewrite_rule @{thms skolem_def [abs_def]})
39949
186a3b447e0b more explicit name
blanchet
parents: 39948
diff changeset
   213
                 RS Global_Theory.get_thm thy "Hilbert_Choice.someI_ex") 1
37617
f73cd4069f69 always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents: 37616
diff changeset
   214
  in
37629
a4f129820562 more elegant cheating
blanchet
parents: 37628
diff changeset
   215
    Goal.prove_internal [ex_tm] conc tacf
a4f129820562 more elegant cheating
blanchet
parents: 37628
diff changeset
   216
    |> forall_intr_list frees
a4f129820562 more elegant cheating
blanchet
parents: 37628
diff changeset
   217
    |> Thm.forall_elim_vars 0  (*Introduce Vars, but don't discharge defs.*)
a4f129820562 more elegant cheating
blanchet
parents: 37628
diff changeset
   218
    |> Thm.varifyT_global
37617
f73cd4069f69 always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents: 37616
diff changeset
   219
  end
24742
73b8b42a36b6 removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents: 24669
diff changeset
   220
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 42333
diff changeset
   221
fun to_definitional_cnf_with_quantifiers ctxt th =
39036
dff91b90d74c use definitional CNFs in Metis rather than plain CNF, following a suggestion by Joe Hurd;
blanchet
parents: 38864
diff changeset
   222
  let
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 42333
diff changeset
   223
    val eqth = cnf.make_cnfx_thm ctxt (HOLogic.dest_Trueprop (prop_of th))
39036
dff91b90d74c use definitional CNFs in Metis rather than plain CNF, following a suggestion by Joe Hurd;
blanchet
parents: 38864
diff changeset
   224
    val eqth = eqth RS @{thm eq_reflection}
dff91b90d74c use definitional CNFs in Metis rather than plain CNF, following a suggestion by Joe Hurd;
blanchet
parents: 38864
diff changeset
   225
    val eqth = eqth RS @{thm TruepropI}
dff91b90d74c use definitional CNFs in Metis rather than plain CNF, following a suggestion by Joe Hurd;
blanchet
parents: 38864
diff changeset
   226
  in Thm.equal_elim eqth th end
dff91b90d74c use definitional CNFs in Metis rather than plain CNF, following a suggestion by Joe Hurd;
blanchet
parents: 38864
diff changeset
   227
39932
acde1b606b0e reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
blanchet
parents: 39931
diff changeset
   228
fun zapped_var_name ((ax_no, cluster_no), skolem) index_no s =
39896
13b3a2ba9ea7 encode axiom number and cluster number in all zapped quantifiers to help discharging new skolemizer assumptions
blanchet
parents: 39894
diff changeset
   229
  (if skolem then new_skolem_var_prefix else new_nonskolem_var_prefix) ^
39932
acde1b606b0e reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
blanchet
parents: 39931
diff changeset
   230
  "_" ^ string_of_int ax_no ^ "_" ^ string_of_int cluster_no ^ "_" ^
40261
7a02144874f3 more work on new Skolemizer without Hilbert_Choice
blanchet
parents: 40260
diff changeset
   231
  string_of_int index_no ^ "_" ^ Name.desymbolize false s
39896
13b3a2ba9ea7 encode axiom number and cluster number in all zapped quantifiers to help discharging new skolemizer assumptions
blanchet
parents: 39894
diff changeset
   232
39899
608b108ec979 compute quantifier dependency graph in new skolemizer
blanchet
parents: 39897
diff changeset
   233
fun cluster_of_zapped_var_name s =
39932
acde1b606b0e reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
blanchet
parents: 39931
diff changeset
   234
  let val get_int = the o Int.fromString o nth (space_explode "_" s) in
acde1b606b0e reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
blanchet
parents: 39931
diff changeset
   235
    ((get_int 1, (get_int 2, get_int 3)),
acde1b606b0e reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
blanchet
parents: 39931
diff changeset
   236
     String.isPrefix new_skolem_var_prefix s)
acde1b606b0e reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
blanchet
parents: 39931
diff changeset
   237
  end
39897
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39896
diff changeset
   238
40260
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   239
fun rename_bound_vars_to_be_zapped ax_no =
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   240
  let
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   241
    fun aux (cluster as (cluster_no, cluster_skolem)) index_no pos t =
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   242
      case t of
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   243
        (t1 as Const (s, _)) $ Abs (s', T, t') =>
39906
blanchet
parents: 39905
diff changeset
   244
        if s = @{const_name all} orelse s = @{const_name All} orelse
blanchet
parents: 39905
diff changeset
   245
           s = @{const_name Ex} then
39932
acde1b606b0e reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
blanchet
parents: 39931
diff changeset
   246
          let
acde1b606b0e reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
blanchet
parents: 39931
diff changeset
   247
            val skolem = (pos = (s = @{const_name Ex}))
acde1b606b0e reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
blanchet
parents: 39931
diff changeset
   248
            val (cluster, index_no) =
acde1b606b0e reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
blanchet
parents: 39931
diff changeset
   249
              if skolem = cluster_skolem then (cluster, index_no)
acde1b606b0e reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
blanchet
parents: 39931
diff changeset
   250
              else ((cluster_no ||> cluster_skolem ? Integer.add 1, skolem), 0)
40260
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   251
            val s' = zapped_var_name cluster index_no s'
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   252
          in t1 $ Abs (s', T, aux cluster (index_no + 1) pos t') end
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   253
        else
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   254
          t
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   255
      | (t1 as Const (s, _)) $ t2 $ t3 =>
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   256
        if s = @{const_name "==>"} orelse s = @{const_name HOL.implies} then
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   257
          t1 $ aux cluster index_no (not pos) t2 $ aux cluster index_no pos t3
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   258
        else if s = @{const_name HOL.conj} orelse
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   259
                s = @{const_name HOL.disj} then
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   260
          t1 $ aux cluster index_no pos t2 $ aux cluster index_no pos t3
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   261
        else
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   262
          t
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   263
      | (t1 as Const (s, _)) $ t2 =>
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   264
        if s = @{const_name Trueprop} then
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   265
          t1 $ aux cluster index_no pos t2
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   266
        else if s = @{const_name Not} then
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   267
          t1 $ aux cluster index_no (not pos) t2
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   268
        else
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   269
          t
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   270
      | _ => t
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   271
  in aux ((ax_no, 0), true) 0 true end
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   272
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   273
fun zap pos ct =
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   274
  ct
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   275
  |> (case term_of ct of
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   276
        Const (s, _) $ Abs (s', _, _) =>
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   277
        if s = @{const_name all} orelse s = @{const_name All} orelse
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   278
           s = @{const_name Ex} then
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   279
          Thm.dest_comb #> snd #> Thm.dest_abs (SOME s') #> snd #> zap pos
39906
blanchet
parents: 39905
diff changeset
   280
        else
blanchet
parents: 39905
diff changeset
   281
          Conv.all_conv
blanchet
parents: 39905
diff changeset
   282
      | Const (s, _) $ _ $ _ =>
blanchet
parents: 39905
diff changeset
   283
        if s = @{const_name "==>"} orelse s = @{const_name implies} then
40260
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   284
          Conv.combination_conv (Conv.arg_conv (zap (not pos))) (zap pos)
39906
blanchet
parents: 39905
diff changeset
   285
        else if s = @{const_name conj} orelse s = @{const_name disj} then
40260
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   286
          Conv.combination_conv (Conv.arg_conv (zap pos)) (zap pos)
39906
blanchet
parents: 39905
diff changeset
   287
        else
blanchet
parents: 39905
diff changeset
   288
          Conv.all_conv
blanchet
parents: 39905
diff changeset
   289
      | Const (s, _) $ _ =>
40260
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   290
        if s = @{const_name Trueprop} then Conv.arg_conv (zap pos)
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   291
        else if s = @{const_name Not} then Conv.arg_conv (zap (not pos))
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   292
        else Conv.all_conv
39906
blanchet
parents: 39905
diff changeset
   293
      | _ => Conv.all_conv)
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   294
41225
bd4ecd48c21f refer to regular structure Simplifier;
wenzelm
parents: 40263
diff changeset
   295
fun ss_only ths = Simplifier.clear_ss HOL_basic_ss addsimps ths
39901
75d792edf634 make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents: 39900
diff changeset
   296
40261
7a02144874f3 more work on new Skolemizer without Hilbert_Choice
blanchet
parents: 40260
diff changeset
   297
val cheat_choice =
39901
75d792edf634 make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents: 39900
diff changeset
   298
  @{prop "ALL x. EX y. Q x y ==> EX f. ALL x. Q x (f x)"}
75d792edf634 make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents: 39900
diff changeset
   299
  |> Logic.varify_global
75d792edf634 make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents: 39900
diff changeset
   300
  |> Skip_Proof.make_thm @{theory}
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   301
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   302
(* Converts an Isabelle theorem into NNF. *)
39901
75d792edf634 make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents: 39900
diff changeset
   303
fun nnf_axiom choice_ths new_skolemizer ax_no th ctxt =
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   304
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42351
diff changeset
   305
    val thy = Proof_Context.theory_of ctxt
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   306
    val th =
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   307
      th |> transform_elim_theorem
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   308
         |> zero_var_indexes
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   309
         |> new_skolemizer ? forall_intr_vars
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   310
    val (th, ctxt) = Variable.import true [th] ctxt |>> snd |>> the_single
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   311
    val th = th |> Conv.fconv_rule Object_Logic.atomize
47954
aada9fd08b58 make higher-order goals more first-order via extensionality
blanchet
parents: 47953
diff changeset
   312
                |> cong_extensionalize_thm thy
aada9fd08b58 make higher-order goals more first-order via extensionality
blanchet
parents: 47953
diff changeset
   313
                |> abs_extensionalize_thm ctxt
aada9fd08b58 make higher-order goals more first-order via extensionality
blanchet
parents: 47953
diff changeset
   314
                |> make_nnf ctxt
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   315
  in
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   316
    if new_skolemizer then
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   317
      let
39901
75d792edf634 make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents: 39900
diff changeset
   318
        fun skolemize choice_ths =
39950
f3c4849868b8 got rid of overkill "meson_choice" attribute;
blanchet
parents: 39949
diff changeset
   319
          skolemize_with_choice_theorems ctxt choice_ths
39901
75d792edf634 make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents: 39900
diff changeset
   320
          #> simplify (ss_only @{thms all_simps[symmetric]})
42347
53e444ecb525 properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
blanchet
parents: 42336
diff changeset
   321
        val no_choice = null choice_ths
39901
75d792edf634 make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents: 39900
diff changeset
   322
        val pull_out =
42347
53e444ecb525 properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
blanchet
parents: 42336
diff changeset
   323
          if no_choice then
53e444ecb525 properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
blanchet
parents: 42336
diff changeset
   324
            simplify (ss_only @{thms all_simps[symmetric] ex_simps[symmetric]})
53e444ecb525 properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
blanchet
parents: 42336
diff changeset
   325
          else
53e444ecb525 properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
blanchet
parents: 42336
diff changeset
   326
            skolemize choice_ths
53e444ecb525 properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
blanchet
parents: 42336
diff changeset
   327
        val discharger_th = th |> pull_out
40260
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   328
        val discharger_th =
42347
53e444ecb525 properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
blanchet
parents: 42336
diff changeset
   329
          discharger_th |> has_too_many_clauses ctxt (concl_of discharger_th)
53e444ecb525 properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
blanchet
parents: 42336
diff changeset
   330
                           ? (to_definitional_cnf_with_quantifiers ctxt
53e444ecb525 properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
blanchet
parents: 42336
diff changeset
   331
                              #> pull_out)
42099
447fa058ab22 avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents: 42098
diff changeset
   332
        val zapped_th =
40263
51ed7a5ad4c5 restructure Skolemization code slightly
blanchet
parents: 40262
diff changeset
   333
          discharger_th |> prop_of |> rename_bound_vars_to_be_zapped ax_no
51ed7a5ad4c5 restructure Skolemization code slightly
blanchet
parents: 40262
diff changeset
   334
          |> (if no_choice then
51ed7a5ad4c5 restructure Skolemization code slightly
blanchet
parents: 40262
diff changeset
   335
                Skip_Proof.make_thm thy #> skolemize [cheat_choice] #> cprop_of
51ed7a5ad4c5 restructure Skolemization code slightly
blanchet
parents: 40262
diff changeset
   336
              else
51ed7a5ad4c5 restructure Skolemization code slightly
blanchet
parents: 40262
diff changeset
   337
                cterm_of thy)
42099
447fa058ab22 avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents: 42098
diff changeset
   338
          |> zap true
447fa058ab22 avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents: 42098
diff changeset
   339
        val fixes =
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 42333
diff changeset
   340
          [] |> Term.add_free_names (prop_of zapped_th)
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 42333
diff changeset
   341
             |> filter is_zapped_var_name
42269
554e90f9db0c tuned comment
blanchet
parents: 42122
diff changeset
   342
        val ctxt' = ctxt |> Variable.add_fixes_direct fixes
42099
447fa058ab22 avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents: 42098
diff changeset
   343
        val fully_skolemized_t =
42333
23bb0784b5d0 try to repair out-of-sync situations in Metis
blanchet
parents: 42269
diff changeset
   344
          zapped_th |> singleton (Variable.export ctxt' ctxt)
23bb0784b5d0 try to repair out-of-sync situations in Metis
blanchet
parents: 42269
diff changeset
   345
                    |> cprop_of |> Thm.dest_equals |> snd |> term_of
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   346
      in
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   347
        if exists_subterm (fn Var ((s, _), _) =>
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   348
                              String.isPrefix new_skolem_var_prefix s
40260
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   349
                            | _ => false) fully_skolemized_t then
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   350
          let
40260
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   351
            val (fully_skolemized_ct, ctxt) =
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   352
              Variable.import_terms true [fully_skolemized_t] ctxt
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   353
              |>> the_single |>> cterm_of thy
40260
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   354
          in
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   355
            (SOME (discharger_th, fully_skolemized_ct),
40262
8403085384eb ensure that MESON correctly preserves the name of variables (needed by the new Skolemizer)
blanchet
parents: 40261
diff changeset
   356
             (Thm.assume fully_skolemized_ct, ctxt))
40260
73ecbe2d432b fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents: 39962
diff changeset
   357
          end
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   358
       else
40262
8403085384eb ensure that MESON correctly preserves the name of variables (needed by the new Skolemizer)
blanchet
parents: 40261
diff changeset
   359
         (NONE, (th, ctxt))
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   360
      end
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   361
    else
42347
53e444ecb525 properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
blanchet
parents: 42336
diff changeset
   362
      (NONE, (th |> has_too_many_clauses ctxt (concl_of th)
53e444ecb525 properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
blanchet
parents: 42336
diff changeset
   363
                    ? to_definitional_cnf_with_quantifiers ctxt, ctxt))
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   364
  end
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   365
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   366
(* Convert a theorem to CNF, with additional premises due to skolemization. *)
45508
b216dc1b3630 started implementing lambda-lifting in Metis
blanchet
parents: 44241
diff changeset
   367
fun cnf_axiom ctxt0 new_skolemizer combinators ax_no th =
37626
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37620
diff changeset
   368
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42351
diff changeset
   369
    val thy = Proof_Context.theory_of ctxt0
39950
f3c4849868b8 got rid of overkill "meson_choice" attribute;
blanchet
parents: 39949
diff changeset
   370
    val choice_ths = choice_theorems thy
40262
8403085384eb ensure that MESON correctly preserves the name of variables (needed by the new Skolemizer)
blanchet
parents: 40261
diff changeset
   371
    val (opt, (nnf_th, ctxt)) =
8403085384eb ensure that MESON correctly preserves the name of variables (needed by the new Skolemizer)
blanchet
parents: 40261
diff changeset
   372
      nnf_axiom choice_ths new_skolemizer ax_no th ctxt0
39894
35ae5cf8c96a encode number of skolem assumptions in them, for more efficient retrieval later
blanchet
parents: 39891
diff changeset
   373
    fun clausify th =
42347
53e444ecb525 properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
blanchet
parents: 42336
diff changeset
   374
      make_cnf (if new_skolemizer orelse null choice_ths then []
53e444ecb525 properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
blanchet
parents: 42336
diff changeset
   375
                else map (old_skolem_theorem_from_def thy) (old_skolem_defs th))
43964
9338aa218f09 thread proper context through, to make sure that "using [[meson_max_clauses = 200]]" is not ignored when clausifying the conjecture
blanchet
parents: 43324
diff changeset
   376
               th ctxt ctxt
42347
53e444ecb525 properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
blanchet
parents: 42336
diff changeset
   377
    val (cnf_ths, ctxt) = clausify nnf_th
39894
35ae5cf8c96a encode number of skolem assumptions in them, for more efficient retrieval later
blanchet
parents: 39891
diff changeset
   378
    fun intr_imp ct th =
39950
f3c4849868b8 got rid of overkill "meson_choice" attribute;
blanchet
parents: 39949
diff changeset
   379
      Thm.instantiate ([], map (pairself (cterm_of thy))
39962
d42ddd7407ca qualify "Meson." and "Metis." constants so that the ML files can be reloaded later (e.g. for testing/development)
blanchet
parents: 39950
diff changeset
   380
                               [(Var (("i", 0), @{typ nat}),
39902
bb43fe4fac93 avoid dependency on "int"
blanchet
parents: 39901
diff changeset
   381
                                 HOLogic.mk_nat ax_no)])
39962
d42ddd7407ca qualify "Meson." and "Metis." constants so that the ML files can be reloaded later (e.g. for testing/development)
blanchet
parents: 39950
diff changeset
   382
                      (zero_var_indexes @{thm skolem_COMBK_D})
39894
35ae5cf8c96a encode number of skolem assumptions in them, for more efficient retrieval later
blanchet
parents: 39891
diff changeset
   383
      RS Thm.implies_intr ct th
37626
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37620
diff changeset
   384
  in
39897
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39896
diff changeset
   385
    (opt |> Option.map (I #>> singleton (Variable.export ctxt ctxt0)
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39896
diff changeset
   386
                        ##> (term_of #> HOLogic.dest_Trueprop
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39896
diff changeset
   387
                             #> singleton (Variable.export_terms ctxt ctxt0))),
45508
b216dc1b3630 started implementing lambda-lifting in Metis
blanchet
parents: 44241
diff changeset
   388
     cnf_ths |> map (combinators ? introduce_combinators_in_theorem
39894
35ae5cf8c96a encode number of skolem assumptions in them, for more efficient retrieval later
blanchet
parents: 39891
diff changeset
   389
                     #> (case opt of SOME (_, ct) => intr_imp ct | NONE => I))
39897
e26d5344e1b7 compute substitutions in new skolemizer
blanchet
parents: 39896
diff changeset
   390
             |> Variable.export ctxt ctxt0
39950
f3c4849868b8 got rid of overkill "meson_choice" attribute;
blanchet
parents: 39949
diff changeset
   391
             |> finish_cnf
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   392
             |> map Thm.close_derivation)
37626
1146291fe718 move blacklisting completely out of the clausifier;
blanchet
parents: 37620
diff changeset
   393
  end
39887
74939e2afb95 second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents: 39886
diff changeset
   394
  handle THM _ => (NONE, [])
27184
b1483d423512 export just one setup function;
wenzelm
parents: 27179
diff changeset
   395
20461
d689ad772b2c skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents: 20457
diff changeset
   396
end;