src/HOL/Tools/Sledgehammer/sledgehammer_fact_preprocessor.ML
author blanchet
Mon, 22 Mar 2010 10:25:44 +0100
changeset 35870 05f3af00cc7e
parent 35845 e5980f0ad025
parent 35869 cac366550624
child 35961 943e2582dc87
permissions -rw-r--r--
merged
blanchet@35826
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_fact_preprocessor.ML
wenzelm@33319
     2
    Author:     Jia Meng, Cambridge University Computer Laboratory
paulson@15347
     3
wenzelm@20461
     4
Transformation of axiom rules (elim/intro/etc) into CNF forms.
paulson@15347
     5
*)
paulson@15347
     6
blanchet@35826
     7
signature SLEDGEHAMMER_FACT_PREPROCESSOR =
wenzelm@21505
     8
sig
wenzelm@32955
     9
  val trace: bool Unsynchronized.ref
wenzelm@32955
    10
  val trace_msg: (unit -> string) -> unit
blanchet@35865
    11
  val skolem_prefix: string
wenzelm@27179
    12
  val cnf_axiom: theory -> thm -> thm list
wenzelm@24669
    13
  val pairname: thm -> string * thm
wenzelm@27184
    14
  val multi_base_blacklist: string list
paulson@25243
    15
  val bad_for_atp: thm -> bool
wenzelm@35568
    16
  val type_has_topsort: typ -> bool
wenzelm@27179
    17
  val cnf_rules_pairs: theory -> (string * thm) list -> (thm * (string * int)) list
wenzelm@24669
    18
  val neg_clausify: thm list -> thm list
paulson@24827
    19
  val combinators: thm -> thm
wenzelm@32261
    20
  val neg_conjecture_clauses: Proof.context -> thm -> int -> thm list * (string * typ) list
wenzelm@32740
    21
  val suppress_endtheory: bool Unsynchronized.ref
wenzelm@32740
    22
    (*for emergency use where endtheory causes problems*)
wenzelm@24669
    23
  val setup: theory -> theory
wenzelm@21505
    24
end;
mengj@19196
    25
blanchet@35826
    26
structure Sledgehammer_Fact_Preprocessor : SLEDGEHAMMER_FACT_PREPROCESSOR =
paulson@15997
    27
struct
paulson@15347
    28
blanchet@35865
    29
open Sledgehammer_FOL_Clause
blanchet@35865
    30
wenzelm@32955
    31
val trace = Unsynchronized.ref false;
blanchet@35865
    32
fun trace_msg msg = if !trace then tracing (msg ()) else ();
blanchet@35865
    33
blanchet@35865
    34
val skolem_prefix = "sko_"
wenzelm@32955
    35
wenzelm@33832
    36
fun freeze_thm th = #1 (Drule.legacy_freeze_thaw th);
paulson@20863
    37
wenzelm@35568
    38
val type_has_topsort = Term.exists_subtype
wenzelm@35568
    39
  (fn TFree (_, []) => true
wenzelm@35568
    40
    | TVar (_, []) => true
wenzelm@35568
    41
    | _ => false);
wenzelm@27184
    42
wenzelm@28544
    43
paulson@15997
    44
(**** Transformation of Elimination Rules into First-Order Formulas****)
paulson@15347
    45
wenzelm@29064
    46
val cfalse = cterm_of @{theory HOL} HOLogic.false_const;
wenzelm@29064
    47
val ctp_false = cterm_of @{theory HOL} (HOLogic.mk_Trueprop HOLogic.false_const);
paulson@15347
    48
paulson@21430
    49
(*Converts an elim-rule into an equivalent theorem that does not have the
paulson@21430
    50
  predicate variable.  Leaves other theorems unchanged.  We simply instantiate the
paulson@21430
    51
  conclusion variable to False.*)
paulson@16009
    52
fun transform_elim th =
paulson@21430
    53
  case concl_of th of    (*conclusion variable*)
wenzelm@24669
    54
       Const("Trueprop",_) $ (v as Var(_,Type("bool",[]))) =>
wenzelm@29064
    55
           Thm.instantiate ([], [(cterm_of @{theory HOL} v, cfalse)]) th
wenzelm@24669
    56
    | v as Var(_, Type("prop",[])) =>
wenzelm@29064
    57
           Thm.instantiate ([], [(cterm_of @{theory HOL} v, ctp_false)]) th
paulson@21430
    58
    | _ => th;
paulson@15997
    59
paulson@24742
    60
(*To enforce single-threading*)
paulson@24742
    61
exception Clausify_failure of theory;
wenzelm@20461
    62
wenzelm@28544
    63
paulson@16009
    64
(**** SKOLEMIZATION BY INFERENCE (lcp) ****)
paulson@16009
    65
paulson@24742
    66
fun rhs_extra_types lhsT rhs =
paulson@24742
    67
  let val lhs_vars = Term.add_tfreesT lhsT []
paulson@24742
    68
      fun add_new_TFrees (TFree v) =
wenzelm@24821
    69
            if member (op =) lhs_vars v then I else insert (op =) (TFree v)
wenzelm@24821
    70
        | add_new_TFrees _ = I
paulson@24742
    71
      val rhs_consts = fold_aterms (fn Const c => insert (op =) c | _ => I) rhs []
paulson@24742
    72
  in fold (#2 #> Term.fold_atyps add_new_TFrees) rhs_consts [] end;
paulson@24742
    73
paulson@18141
    74
(*Traverse a theorem, declaring Skolem function definitions. String s is the suggested
wenzelm@27174
    75
  prefix for the Skolem constant.*)
wenzelm@27174
    76
fun declare_skofuns s th =
wenzelm@27174
    77
  let
wenzelm@33233
    78
    val nref = Unsynchronized.ref 0    (* FIXME ??? *)
wenzelm@27174
    79
    fun dec_sko (Const ("Ex",_) $ (xtp as Abs (_, T, p))) (axs, thy) =
wenzelm@27174
    80
          (*Existential: declare a Skolem function, then insert into body and continue*)
wenzelm@27174
    81
          let
blanchet@35865
    82
            val cname = skolem_prefix ^ s ^ "_" ^ Int.toString (Unsynchronized.inc nref)
wenzelm@29265
    83
            val args0 = OldTerm.term_frees xtp  (*get the formal parameter list*)
wenzelm@27174
    84
            val Ts = map type_of args0
wenzelm@27174
    85
            val extraTs = rhs_extra_types (Ts ---> T) xtp
wenzelm@27174
    86
            val argsx = map (fn T => Free (gensym "vsk", T)) extraTs
wenzelm@27174
    87
            val args = argsx @ args0
wenzelm@27174
    88
            val cT = extraTs ---> Ts ---> T
wenzelm@27174
    89
            val rhs = list_abs_free (map dest_Free args, HOLogic.choice_const T $ xtp)
wenzelm@27174
    90
                    (*Forms a lambda-abstraction over the formal parameters*)
wenzelm@28110
    91
            val (c, thy') =
wenzelm@33173
    92
              Sign.declare_const ((Binding.conceal (Binding.name cname), cT), NoSyn) thy
wenzelm@27174
    93
            val cdef = cname ^ "_def"
wenzelm@33233
    94
            val thy'' =
wenzelm@33233
    95
              Theory.add_defs_i true false [(Binding.name cdef, Logic.mk_equals (c, rhs))] thy'
haftmann@28965
    96
            val ax = Thm.axiom thy'' (Sign.full_bname thy'' cdef)
wenzelm@27174
    97
          in dec_sko (subst_bound (list_comb (c, args), p)) (ax :: axs, thy'') end
wenzelm@32994
    98
      | dec_sko (Const ("All", _) $ (Abs (a, T, p))) thx =
wenzelm@27174
    99
          (*Universal quant: insert a free variable into body and continue*)
wenzelm@29270
   100
          let val fname = Name.variant (OldTerm.add_term_names (p, [])) a
wenzelm@27174
   101
          in dec_sko (subst_bound (Free (fname, T), p)) thx end
wenzelm@27174
   102
      | dec_sko (Const ("op &", _) $ p $ q) thx = dec_sko q (dec_sko p thx)
wenzelm@27174
   103
      | dec_sko (Const ("op |", _) $ p $ q) thx = dec_sko q (dec_sko p thx)
wenzelm@27174
   104
      | dec_sko (Const ("Trueprop", _) $ p) thx = dec_sko p thx
wenzelm@27174
   105
      | dec_sko t thx = thx (*Do nothing otherwise*)
wenzelm@27174
   106
  in fn thy => dec_sko (Thm.prop_of th) ([], thy) end;
paulson@18141
   107
paulson@18141
   108
(*Traverse a theorem, accumulating Skolem function definitions.*)
paulson@22731
   109
fun assume_skofuns s th =
wenzelm@33233
   110
  let val sko_count = Unsynchronized.ref 0   (* FIXME ??? *)
paulson@22731
   111
      fun dec_sko (Const ("Ex",_) $ (xtp as Abs(_,T,p))) defs =
wenzelm@20461
   112
            (*Existential: declare a Skolem function, then insert into body and continue*)
wenzelm@20461
   113
            let val skos = map (#1 o Logic.dest_equals) defs  (*existing sko fns*)
haftmann@33040
   114
                val args = subtract (op =) skos (OldTerm.term_frees xtp) (*the formal parameters*)
wenzelm@20461
   115
                val Ts = map type_of args
wenzelm@20461
   116
                val cT = Ts ---> T
blanchet@35865
   117
                val id = skolem_prefix ^ s ^ "_" ^ Int.toString (Unsynchronized.inc sko_count)
paulson@22731
   118
                val c = Free (id, cT)
wenzelm@20461
   119
                val rhs = list_abs_free (map dest_Free args,
wenzelm@20461
   120
                                         HOLogic.choice_const T $ xtp)
wenzelm@20461
   121
                      (*Forms a lambda-abstraction over the formal parameters*)
wenzelm@27330
   122
                val def = Logic.mk_equals (c, rhs)
wenzelm@20461
   123
            in dec_sko (subst_bound (list_comb(c,args), p))
wenzelm@20461
   124
                       (def :: defs)
wenzelm@20461
   125
            end
wenzelm@32994
   126
        | dec_sko (Const ("All",_) $ Abs (a, T, p)) defs =
wenzelm@20461
   127
            (*Universal quant: insert a free variable into body and continue*)
wenzelm@29270
   128
            let val fname = Name.variant (OldTerm.add_term_names (p,[])) a
wenzelm@20461
   129
            in dec_sko (subst_bound (Free(fname,T), p)) defs end
wenzelm@20461
   130
        | dec_sko (Const ("op &", _) $ p $ q) defs = dec_sko q (dec_sko p defs)
wenzelm@20461
   131
        | dec_sko (Const ("op |", _) $ p $ q) defs = dec_sko q (dec_sko p defs)
wenzelm@20461
   132
        | dec_sko (Const ("Trueprop", _) $ p) defs = dec_sko p defs
wenzelm@20461
   133
        | dec_sko t defs = defs (*Do nothing otherwise*)
paulson@20419
   134
  in  dec_sko (prop_of th) []  end;
paulson@20419
   135
paulson@20419
   136
paulson@24827
   137
(**** REPLACING ABSTRACTIONS BY COMBINATORS ****)
paulson@20419
   138
paulson@20419
   139
(*Returns the vars of a theorem*)
paulson@20419
   140
fun vars_of_thm th =
wenzelm@22691
   141
  map (Thm.cterm_of (theory_of_thm th) o Var) (Thm.fold_terms Term.add_vars th []);
paulson@20419
   142
paulson@20419
   143
(*Make a version of fun_cong with a given variable name*)
paulson@20419
   144
local
paulson@20419
   145
    val fun_cong' = fun_cong RS asm_rl; (*renumber f, g to prevent clashes with (a,0)*)
paulson@20419
   146
    val cx = hd (vars_of_thm fun_cong');
paulson@20419
   147
    val ty = typ_of (ctyp_of_term cx);
paulson@20445
   148
    val thy = theory_of_thm fun_cong;
paulson@20419
   149
    fun mkvar a = cterm_of thy (Var((a,0),ty));
paulson@20419
   150
in
paulson@20419
   151
fun xfun_cong x = Thm.instantiate ([], [(cx, mkvar x)]) fun_cong'
paulson@20419
   152
end;
paulson@20419
   153
paulson@20863
   154
(*Removes the lambdas from an equation of the form t = (%x. u).  A non-negative n,
paulson@20863
   155
  serves as an upper bound on how many to remove.*)
paulson@20863
   156
fun strip_lambdas 0 th = th
wenzelm@24669
   157
  | strip_lambdas n th =
paulson@20863
   158
      case prop_of th of
wenzelm@24669
   159
          _ $ (Const ("op =", _) $ _ $ Abs (x,_,_)) =>
wenzelm@24669
   160
              strip_lambdas (n-1) (freeze_thm (th RS xfun_cong x))
wenzelm@24669
   161
        | _ => th;
paulson@20419
   162
wenzelm@24669
   163
val lambda_free = not o Term.has_abs;
wenzelm@20461
   164
wenzelm@32011
   165
val [f_B,g_B] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_B}));
wenzelm@32011
   166
val [g_C,f_C] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_C}));
wenzelm@32011
   167
val [f_S,g_S] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_S}));
paulson@20710
   168
paulson@24827
   169
(*FIXME: requires more use of cterm constructors*)
paulson@24827
   170
fun abstract ct =
wenzelm@28544
   171
  let
wenzelm@28544
   172
      val thy = theory_of_cterm ct
paulson@25256
   173
      val Abs(x,_,body) = term_of ct
paulson@24827
   174
      val Type("fun",[xT,bodyT]) = typ_of (ctyp_of_term ct)
paulson@24827
   175
      val cxT = ctyp_of thy xT and cbodyT = ctyp_of thy bodyT
wenzelm@27184
   176
      fun makeK() = instantiate' [SOME cxT, SOME cbodyT] [SOME (cterm_of thy body)] @{thm abs_K}
paulson@24827
   177
  in
paulson@24827
   178
      case body of
paulson@24827
   179
          Const _ => makeK()
paulson@24827
   180
        | Free _ => makeK()
paulson@24827
   181
        | Var _ => makeK()  (*though Var isn't expected*)
wenzelm@27184
   182
        | Bound 0 => instantiate' [SOME cxT] [] @{thm abs_I} (*identity: I*)
paulson@24827
   183
        | rator$rand =>
wenzelm@27184
   184
            if loose_bvar1 (rator,0) then (*C or S*)
wenzelm@27179
   185
               if loose_bvar1 (rand,0) then (*S*)
wenzelm@27179
   186
                 let val crator = cterm_of thy (Abs(x,xT,rator))
wenzelm@27179
   187
                     val crand = cterm_of thy (Abs(x,xT,rand))
wenzelm@27184
   188
                     val abs_S' = cterm_instantiate [(f_S,crator),(g_S,crand)] @{thm abs_S}
wenzelm@27184
   189
                     val (_,rhs) = Thm.dest_equals (cprop_of abs_S')
wenzelm@27179
   190
                 in
wenzelm@27179
   191
                   Thm.transitive abs_S' (Conv.binop_conv abstract rhs)
wenzelm@27179
   192
                 end
wenzelm@27179
   193
               else (*C*)
wenzelm@27179
   194
                 let val crator = cterm_of thy (Abs(x,xT,rator))
wenzelm@27184
   195
                     val abs_C' = cterm_instantiate [(f_C,crator),(g_C,cterm_of thy rand)] @{thm abs_C}
wenzelm@27184
   196
                     val (_,rhs) = Thm.dest_equals (cprop_of abs_C')
wenzelm@27179
   197
                 in
wenzelm@27179
   198
                   Thm.transitive abs_C' (Conv.fun_conv (Conv.arg_conv abstract) rhs)
wenzelm@27179
   199
                 end
wenzelm@27184
   200
            else if loose_bvar1 (rand,0) then (*B or eta*)
wenzelm@27179
   201
               if rand = Bound 0 then eta_conversion ct
wenzelm@27179
   202
               else (*B*)
wenzelm@27179
   203
                 let val crand = cterm_of thy (Abs(x,xT,rand))
wenzelm@27179
   204
                     val crator = cterm_of thy rator
wenzelm@27184
   205
                     val abs_B' = cterm_instantiate [(f_B,crator),(g_B,crand)] @{thm abs_B}
wenzelm@27184
   206
                     val (_,rhs) = Thm.dest_equals (cprop_of abs_B')
wenzelm@27179
   207
                 in
wenzelm@27179
   208
                   Thm.transitive abs_B' (Conv.arg_conv abstract rhs)
wenzelm@27179
   209
                 end
wenzelm@27179
   210
            else makeK()
paulson@24827
   211
        | _ => error "abstract: Bad term"
paulson@24827
   212
  end;
paulson@20863
   213
paulson@20419
   214
(*Traverse a theorem, declaring abstraction function definitions. String s is the suggested
wenzelm@28544
   215
  prefix for the constants.*)
paulson@24827
   216
fun combinators_aux ct =
paulson@24827
   217
  if lambda_free (term_of ct) then reflexive ct
paulson@24827
   218
  else
paulson@24827
   219
  case term_of ct of
paulson@24827
   220
      Abs _ =>
wenzelm@32994
   221
        let val (cv, cta) = Thm.dest_abs NONE ct
wenzelm@32994
   222
            val (v, _) = dest_Free (term_of cv)
wenzelm@27179
   223
            val u_th = combinators_aux cta
wenzelm@27179
   224
            val cu = Thm.rhs_of u_th
wenzelm@27179
   225
            val comb_eq = abstract (Thm.cabs cv cu)
wenzelm@28544
   226
        in transitive (abstract_rule v cv u_th) comb_eq end
wenzelm@32994
   227
    | _ $ _ =>
wenzelm@32994
   228
        let val (ct1, ct2) = Thm.dest_comb ct
wenzelm@27179
   229
        in  combination (combinators_aux ct1) (combinators_aux ct2)  end;
wenzelm@27184
   230
paulson@24827
   231
fun combinators th =
wenzelm@27184
   232
  if lambda_free (prop_of th) then th
paulson@24827
   233
  else
wenzelm@28544
   234
    let val th = Drule.eta_contraction_rule th
wenzelm@27179
   235
        val eqth = combinators_aux (cprop_of th)
paulson@25256
   236
    in  equal_elim eqth th   end
wenzelm@27184
   237
    handle THM (msg,_,_) =>
wenzelm@32111
   238
      (warning (cat_lines
wenzelm@32111
   239
        ["Error in the combinator translation of " ^ Display.string_of_thm_without_context th,
wenzelm@32111
   240
          "  Exception message: " ^ msg]);
paulson@25256
   241
       TrueI);  (*A type variable of sort {} will cause make abstraction fail.*)
paulson@16009
   242
paulson@16009
   243
(*cterms are used throughout for efficiency*)
wenzelm@29064
   244
val cTrueprop = Thm.cterm_of @{theory HOL} HOLogic.Trueprop;
paulson@16009
   245
paulson@16009
   246
(*cterm version of mk_cTrueprop*)
paulson@16009
   247
fun c_mkTrueprop A = Thm.capply cTrueprop A;
paulson@16009
   248
paulson@16009
   249
(*Given an abstraction over n variables, replace the bound variables by free
paulson@16009
   250
  ones. Return the body, along with the list of free variables.*)
wenzelm@20461
   251
fun c_variant_abs_multi (ct0, vars) =
paulson@16009
   252
      let val (cv,ct) = Thm.dest_abs NONE ct0
paulson@16009
   253
      in  c_variant_abs_multi (ct, cv::vars)  end
paulson@16009
   254
      handle CTERM _ => (ct0, rev vars);
paulson@16009
   255
wenzelm@20461
   256
(*Given the definition of a Skolem function, return a theorem to replace
wenzelm@20461
   257
  an existential formula by a use of that function.
paulson@18141
   258
   Example: "EX x. x : A & x ~: B ==> sko A B : A & sko A B ~: B"  [.] *)
wenzelm@20461
   259
fun skolem_of_def def =
wenzelm@22902
   260
  let val (c,rhs) = Thm.dest_equals (cprop_of (freeze_thm def))
paulson@16009
   261
      val (ch, frees) = c_variant_abs_multi (rhs, [])
paulson@18141
   262
      val (chilbert,cabs) = Thm.dest_comb ch
wenzelm@26627
   263
      val thy = Thm.theory_of_cterm chilbert
wenzelm@26627
   264
      val t = Thm.term_of chilbert
paulson@18141
   265
      val T = case t of Const ("Hilbert_Choice.Eps", Type("fun",[_,T])) => T
paulson@18141
   266
                      | _ => raise THM ("skolem_of_def: expected Eps", 0, [def])
wenzelm@22596
   267
      val cex = Thm.cterm_of thy (HOLogic.exists_const T)
paulson@16009
   268
      val ex_tm = c_mkTrueprop (Thm.capply cex cabs)
paulson@16009
   269
      and conc =  c_mkTrueprop (Drule.beta_conv cabs (Drule.list_comb(c,frees)));
haftmann@31454
   270
      fun tacf [prem] = rewrite_goals_tac [def] THEN rtac (prem RS @{thm someI_ex}) 1
wenzelm@23352
   271
  in  Goal.prove_internal [ex_tm] conc tacf
paulson@18141
   272
       |> forall_intr_list frees
wenzelm@26653
   273
       |> Thm.forall_elim_vars 0  (*Introduce Vars, but don't discharge defs.*)
wenzelm@35845
   274
       |> Thm.varifyT_global
paulson@18141
   275
  end;
paulson@16009
   276
paulson@24742
   277
paulson@20863
   278
(*Converts an Isabelle theorem (intro, elim or simp format, even higher-order) into NNF.*)
paulson@24937
   279
fun to_nnf th ctxt0 =
wenzelm@27179
   280
  let val th1 = th |> transform_elim |> zero_var_indexes
wenzelm@32274
   281
      val ((_, [th2]), ctxt) = Variable.import true [th1] ctxt0
wenzelm@32274
   282
      val th3 = th2
wenzelm@35625
   283
        |> Conv.fconv_rule Object_Logic.atomize
wenzelm@32274
   284
        |> Meson.make_nnf ctxt |> strip_lambdas ~1
paulson@24937
   285
  in  (th3, ctxt)  end;
paulson@16009
   286
paulson@18141
   287
(*Generate Skolem functions for a theorem supplied in nnf*)
paulson@24937
   288
fun assume_skolem_of_def s th =
paulson@22731
   289
  map (skolem_of_def o assume o (cterm_of (theory_of_thm th))) (assume_skofuns s th);
paulson@18141
   290
paulson@25007
   291
blanchet@35826
   292
(*** Blacklisting (duplicated in "Sledgehammer_Fact_Filter"?) ***)
paulson@25007
   293
paulson@25007
   294
val max_lambda_nesting = 3;
wenzelm@27184
   295
paulson@25007
   296
fun excessive_lambdas (f$t, k) = excessive_lambdas (f,k) orelse excessive_lambdas (t,k)
paulson@25007
   297
  | excessive_lambdas (Abs(_,_,t), k) = k=0 orelse excessive_lambdas (t,k-1)
paulson@25007
   298
  | excessive_lambdas _ = false;
paulson@25007
   299
paulson@25007
   300
fun is_formula_type T = (T = HOLogic.boolT orelse T = propT);
paulson@25007
   301
paulson@25007
   302
(*Don't count nested lambdas at the level of formulas, as they are quantifiers*)
paulson@25007
   303
fun excessive_lambdas_fm Ts (Abs(_,T,t)) = excessive_lambdas_fm (T::Ts) t
paulson@25007
   304
  | excessive_lambdas_fm Ts t =
paulson@25007
   305
      if is_formula_type (fastype_of1 (Ts, t))
paulson@25007
   306
      then exists (excessive_lambdas_fm Ts) (#2 (strip_comb t))
paulson@25007
   307
      else excessive_lambdas (t, max_lambda_nesting);
paulson@25007
   308
wenzelm@33027
   309
(*The max apply_depth of any metis call in Metis_Examples (on 31-10-2007) was 11.*)
paulson@25256
   310
val max_apply_depth = 15;
wenzelm@27184
   311
paulson@25256
   312
fun apply_depth (f$t) = Int.max (apply_depth f, apply_depth t + 1)
paulson@25256
   313
  | apply_depth (Abs(_,_,t)) = apply_depth t
paulson@25256
   314
  | apply_depth _ = 0;
paulson@25256
   315
wenzelm@27184
   316
fun too_complex t =
wenzelm@27184
   317
  apply_depth t > max_apply_depth orelse
paulson@26562
   318
  Meson.too_many_clauses NONE t orelse
paulson@25256
   319
  excessive_lambdas_fm [] t;
wenzelm@27184
   320
paulson@25243
   321
fun is_strange_thm th =
paulson@25243
   322
  case head_of (concl_of th) of
wenzelm@33314
   323
      Const (a, _) => (a <> "Trueprop" andalso a <> "==")
paulson@25243
   324
    | _ => false;
paulson@25243
   325
wenzelm@27184
   326
fun bad_for_atp th =
wenzelm@33314
   327
  too_complex (prop_of th)
wenzelm@35568
   328
  orelse exists_type type_has_topsort (prop_of th)
paulson@25761
   329
  orelse is_strange_thm th;
paulson@25243
   330
paulson@25007
   331
val multi_base_blacklist =
paulson@25256
   332
  ["defs","select_defs","update_defs","induct","inducts","split","splits","split_asm",
blanchet@35828
   333
   "cases","ext_cases"];  (* FIXME put other record thms here, or declare as "no_atp" *)
paulson@25007
   334
paulson@21071
   335
(*Keep the full complexity of the original name*)
wenzelm@30364
   336
fun flatten_name s = space_implode "_X" (Long_Name.explode s);
paulson@21071
   337
paulson@22731
   338
fun fake_name th =
wenzelm@27865
   339
  if Thm.has_name_hint th then flatten_name (Thm.get_name_hint th)
paulson@22731
   340
  else gensym "unknown_thm_";
paulson@22731
   341
wenzelm@27184
   342
(*Skolemize a named theorem, with Skolem functions as additional premises.*)
wenzelm@27184
   343
fun skolem_thm (s, th) =
wenzelm@30364
   344
  if member (op =) multi_base_blacklist (Long_Name.base_name s) orelse bad_for_atp th then []
wenzelm@27184
   345
  else
wenzelm@27184
   346
    let
wenzelm@27184
   347
      val ctxt0 = Variable.thm_context th
wenzelm@27184
   348
      val (nnfth, ctxt1) = to_nnf th ctxt0
wenzelm@27184
   349
      val (cnfs, ctxt2) = Meson.make_cnf (assume_skolem_of_def s nnfth) nnfth ctxt1
wenzelm@27184
   350
    in  cnfs |> map combinators |> Variable.export ctxt2 ctxt0 |> Meson.finish_cnf  end
wenzelm@27184
   351
    handle THM _ => [];
wenzelm@27184
   352
paulson@24742
   353
(*The cache prevents repeated clausification of a theorem, and also repeated declaration of
paulson@24742
   354
  Skolem functions.*)
wenzelm@33522
   355
structure ThmCache = Theory_Data
wenzelm@22846
   356
(
wenzelm@28544
   357
  type T = thm list Thmtab.table * unit Symtab.table;
wenzelm@28544
   358
  val empty = (Thmtab.empty, Symtab.empty);
wenzelm@26618
   359
  val extend = I;
wenzelm@33522
   360
  fun merge ((cache1, seen1), (cache2, seen2)) : T =
wenzelm@27184
   361
    (Thmtab.merge (K true) (cache1, cache2), Symtab.merge (K true) (seen1, seen2));
wenzelm@22846
   362
);
paulson@22516
   363
wenzelm@27184
   364
val lookup_cache = Thmtab.lookup o #1 o ThmCache.get;
wenzelm@27184
   365
val already_seen = Symtab.defined o #2 o ThmCache.get;
wenzelm@20461
   366
wenzelm@27184
   367
val update_cache = ThmCache.map o apfst o Thmtab.update;
wenzelm@27184
   368
fun mark_seen name = ThmCache.map (apsnd (Symtab.update (name, ())));
paulson@25007
   369
wenzelm@20461
   370
(*Exported function to convert Isabelle theorems into axiom clauses*)
wenzelm@27179
   371
fun cnf_axiom thy th0 =
wenzelm@27184
   372
  let val th = Thm.transfer thy th0 in
wenzelm@27184
   373
    case lookup_cache thy th of
wenzelm@27184
   374
      NONE => map Thm.close_derivation (skolem_thm (fake_name th, th))
wenzelm@27184
   375
    | SOME cls => cls
paulson@22516
   376
  end;
paulson@15347
   377
paulson@18141
   378
wenzelm@30291
   379
(**** Rules from the context ****)
paulson@15347
   380
wenzelm@27865
   381
fun pairname th = (Thm.get_name_hint th, th);
wenzelm@27184
   382
paulson@15347
   383
paulson@22471
   384
(**** Translate a set of theorems into CNF ****)
paulson@15347
   385
paulson@19894
   386
fun pair_name_cls k (n, []) = []
paulson@19894
   387
  | pair_name_cls k (n, cls::clss) = (cls, (n,k)) :: pair_name_cls (k+1) (n, clss)
wenzelm@20461
   388
wenzelm@27179
   389
fun cnf_rules_pairs_aux _ pairs [] = pairs
wenzelm@27179
   390
  | cnf_rules_pairs_aux thy pairs ((name,th)::ths) =
wenzelm@27179
   391
      let val pairs' = (pair_name_cls 0 (name, cnf_axiom thy th)) @ pairs
blanchet@35826
   392
                       handle THM _ => pairs |
blanchet@35865
   393
                              CLAUSE _ => pairs
wenzelm@27179
   394
      in  cnf_rules_pairs_aux thy pairs' ths  end;
wenzelm@20461
   395
paulson@21290
   396
(*The combination of rev and tail recursion preserves the original order*)
wenzelm@27179
   397
fun cnf_rules_pairs thy l = cnf_rules_pairs_aux thy [] (rev l);
mengj@19353
   398
mengj@19196
   399
blanchet@35865
   400
(**** Convert all facts of the theory into FOL or HOL clauses ****)
paulson@15347
   401
wenzelm@28544
   402
local
paulson@20457
   403
wenzelm@28544
   404
fun skolem_def (name, th) thy =
wenzelm@28544
   405
  let val ctxt0 = Variable.thm_context th in
wenzelm@28544
   406
    (case try (to_nnf th) ctxt0 of
wenzelm@28544
   407
      NONE => (NONE, thy)
wenzelm@28544
   408
    | SOME (nnfth, ctxt1) =>
wenzelm@28544
   409
        let val (defs, thy') = declare_skofuns (flatten_name name) nnfth thy
wenzelm@28544
   410
        in (SOME (th, ctxt0, ctxt1, nnfth, defs), thy') end)
wenzelm@28544
   411
  end;
wenzelm@28544
   412
wenzelm@28544
   413
fun skolem_cnfs (th, ctxt0, ctxt1, nnfth, defs) =
wenzelm@28544
   414
  let
wenzelm@28544
   415
    val (cnfs, ctxt2) = Meson.make_cnf (map skolem_of_def defs) nnfth ctxt1;
wenzelm@28544
   416
    val cnfs' = cnfs
wenzelm@28544
   417
      |> map combinators
wenzelm@28544
   418
      |> Variable.export ctxt2 ctxt0
wenzelm@28544
   419
      |> Meson.finish_cnf
wenzelm@28544
   420
      |> map Thm.close_derivation;
wenzelm@28544
   421
    in (th, cnfs') end;
wenzelm@28544
   422
wenzelm@28544
   423
in
paulson@24742
   424
wenzelm@27184
   425
fun saturate_skolem_cache thy =
wenzelm@28544
   426
  let
wenzelm@33314
   427
    val facts = PureThy.facts_of thy;
wenzelm@33314
   428
    val new_facts = (facts, []) |-> Facts.fold_static (fn (name, ths) =>
wenzelm@33314
   429
      if Facts.is_concealed facts name orelse already_seen thy name then I
wenzelm@33314
   430
      else cons (name, ths));
wenzelm@28544
   431
    val new_thms = (new_facts, []) |-> fold (fn (name, ths) =>
wenzelm@30364
   432
      if member (op =) multi_base_blacklist (Long_Name.base_name name) then I
wenzelm@28544
   433
      else fold_index (fn (i, th) =>
wenzelm@28544
   434
        if bad_for_atp th orelse is_some (lookup_cache thy th) then I
wenzelm@28544
   435
        else cons (name ^ "_" ^ string_of_int (i + 1), Thm.transfer thy th)) ths);
wenzelm@28544
   436
  in
wenzelm@28544
   437
    if null new_facts then NONE
wenzelm@28544
   438
    else
wenzelm@28544
   439
      let
wenzelm@28544
   440
        val (defs, thy') = thy
wenzelm@28544
   441
          |> fold (mark_seen o #1) new_facts
wenzelm@28544
   442
          |> fold_map skolem_def (sort_distinct (Thm.thm_ord o pairself snd) new_thms)
wenzelm@28544
   443
          |>> map_filter I;
wenzelm@29372
   444
        val cache_entries = Par_List.map skolem_cnfs defs;
wenzelm@28544
   445
      in SOME (fold update_cache cache_entries thy') end
wenzelm@28544
   446
  end;
paulson@24742
   447
wenzelm@28544
   448
end;
paulson@24854
   449
wenzelm@32740
   450
val suppress_endtheory = Unsynchronized.ref false;
wenzelm@27184
   451
wenzelm@27184
   452
fun clause_cache_endtheory thy =
wenzelm@27184
   453
  if ! suppress_endtheory then NONE
wenzelm@27184
   454
  else saturate_skolem_cache thy;
wenzelm@27184
   455
paulson@20457
   456
paulson@22516
   457
(*The cache can be kept smaller by inspecting the prop of each thm. Can ignore all that are
paulson@22516
   458
  lambda_free, but then the individual theory caches become much bigger.*)
paulson@21071
   459
wenzelm@27179
   460
paulson@21999
   461
(*** Converting a subgoal into negated conjecture clauses. ***)
paulson@21999
   462
wenzelm@32274
   463
fun neg_skolemize_tac ctxt =
wenzelm@35625
   464
  EVERY' [rtac ccontr, Object_Logic.atomize_prems_tac, Meson.skolemize_tac ctxt];
paulson@22471
   465
blanchet@35869
   466
val neg_clausify =
blanchet@35869
   467
  Meson.make_clauses_unsorted #> map combinators #> Meson.finish_cnf;
paulson@21999
   468
wenzelm@32261
   469
fun neg_conjecture_clauses ctxt st0 n =
wenzelm@32261
   470
  let
wenzelm@32274
   471
    val st = Seq.hd (neg_skolemize_tac ctxt n st0)
wenzelm@32261
   472
    val ({params, prems, ...}, _) = Subgoal.focus (Variable.set_body false ctxt) n st
wenzelm@32261
   473
  in (neg_clausify prems, map (Term.dest_Free o Thm.term_of o #2) params) end;
paulson@21999
   474
wenzelm@24669
   475
(*Conversion of a subgoal to conjecture clauses. Each clause has
paulson@21999
   476
  leading !!-bound universal variables, to express generality. *)
wenzelm@32261
   477
fun neg_clausify_tac ctxt =
wenzelm@32274
   478
  neg_skolemize_tac ctxt THEN'
wenzelm@32261
   479
  SUBGOAL (fn (prop, i) =>
wenzelm@32261
   480
    let val ts = Logic.strip_assums_hyp prop in
wenzelm@32261
   481
      EVERY'
wenzelm@32286
   482
       [Subgoal.FOCUS
wenzelm@32261
   483
         (fn {prems, ...} =>
wenzelm@32261
   484
           (Method.insert_tac
wenzelm@32261
   485
             (map forall_intr_vars (neg_clausify prems)) i)) ctxt,
wenzelm@32261
   486
        REPEAT_DETERM_N (length ts) o etac thin_rl] i
paulson@21999
   487
     end);
paulson@21999
   488
wenzelm@30722
   489
val neg_clausify_setup =
wenzelm@32261
   490
  Method.setup @{binding neg_clausify} (Scan.succeed (SIMPLE_METHOD' o neg_clausify_tac))
wenzelm@30520
   491
  "conversion of goal to conjecture clauses";
wenzelm@24669
   492
wenzelm@27184
   493
wenzelm@27184
   494
(** Attribute for converting a theorem into clauses **)
wenzelm@27184
   495
wenzelm@30722
   496
val clausify_setup =
wenzelm@30722
   497
  Attrib.setup @{binding clausify}
wenzelm@30722
   498
    (Scan.lift OuterParse.nat >>
wenzelm@30722
   499
      (fn i => Thm.rule_attribute (fn context => fn th =>
wenzelm@30722
   500
          Meson.make_meta_clause (nth (cnf_axiom (Context.theory_of context) th) i))))
wenzelm@30722
   501
  "conversion of theorem to clauses";
wenzelm@27184
   502
wenzelm@27184
   503
wenzelm@27184
   504
(** setup **)
wenzelm@27184
   505
wenzelm@27184
   506
val setup =
wenzelm@30722
   507
  neg_clausify_setup #>
wenzelm@30722
   508
  clausify_setup #>
wenzelm@27184
   509
  perhaps saturate_skolem_cache #>
wenzelm@27184
   510
  Theory.at_end clause_cache_endtheory;
paulson@18510
   511
wenzelm@20461
   512
end;