src/HOL/Tools/Sledgehammer/sledgehammer_hol_clause.ML
author blanchet
Fri, 19 Mar 2010 13:02:18 +0100
changeset 35865 2f8fb5242799
parent 35827 f552152d7747
child 35961 943e2582dc87
permissions -rw-r--r--
more Sledgehammer refactoring
blanchet@35826
     1
(*  Title:      HOL/Sledgehammer/sledgehammer_hol_clause.ML
wenzelm@33319
     2
    Author:     Jia Meng, NICTA
mengj@17998
     3
wenzelm@24311
     4
FOL clauses translated from HOL formulae.
mengj@17998
     5
*)
mengj@17998
     6
blanchet@35826
     7
signature SLEDGEHAMMER_HOL_CLAUSE =
wenzelm@24311
     8
sig
blanchet@35865
     9
  type kind = Sledgehammer_FOL_Clause.kind
blanchet@35865
    10
  type fol_type = Sledgehammer_FOL_Clause.fol_type
blanchet@35865
    11
  type classrel_clause = Sledgehammer_FOL_Clause.classrel_clause
blanchet@35865
    12
  type arity_clause = Sledgehammer_FOL_Clause.arity_clause
wenzelm@24311
    13
  type axiom_name = string
wenzelm@24311
    14
  type polarity = bool
blanchet@35865
    15
  type hol_clause_id = int
blanchet@35865
    16
wenzelm@24311
    17
  datatype combterm =
blanchet@35865
    18
    CombConst of string * fol_type * fol_type list (* Const and Free *) |
blanchet@35865
    19
    CombVar of string * fol_type |
blanchet@35865
    20
    CombApp of combterm * combterm
wenzelm@24311
    21
  datatype literal = Literal of polarity * combterm
blanchet@35865
    22
  datatype hol_clause =
blanchet@35865
    23
    HOLClause of {clause_id: hol_clause_id, axiom_name: axiom_name, th: thm,
blanchet@35865
    24
                  kind: kind, literals: literal list, ctypes_sorts: typ list}
blanchet@35865
    25
blanchet@35865
    26
  val type_of_combterm : combterm -> fol_type
blanchet@35865
    27
  val strip_combterm_comb : combterm -> combterm * combterm list
blanchet@35865
    28
  val literals_of_term : theory -> term -> literal list * typ list
blanchet@35865
    29
  exception TRIVIAL
blanchet@35865
    30
  val make_conjecture_clauses : bool -> theory -> thm list -> hol_clause list
blanchet@35865
    31
  val make_axiom_clauses : bool -> theory ->
blanchet@35865
    32
       (thm * (axiom_name * hol_clause_id)) list -> (axiom_name * hol_clause) list
blanchet@35865
    33
  val get_helper_clauses : bool -> theory -> bool ->
blanchet@35865
    34
       hol_clause list * (thm * (axiom_name * hol_clause_id)) list * string list ->
blanchet@35865
    35
       hol_clause list
blanchet@35865
    36
  val write_tptp_file : bool -> Path.T ->
blanchet@35865
    37
    hol_clause list * hol_clause list * hol_clause list * hol_clause list *
blanchet@35865
    38
    classrel_clause list * arity_clause list ->
immler@31832
    39
    int * int
blanchet@35865
    40
  val write_dfg_file : bool -> Path.T ->
blanchet@35865
    41
    hol_clause list * hol_clause list * hol_clause list * hol_clause list *
blanchet@35865
    42
    classrel_clause list * arity_clause list -> int * int
wenzelm@24311
    43
end
mengj@17998
    44
blanchet@35826
    45
structure Sledgehammer_HOL_Clause : SLEDGEHAMMER_HOL_CLAUSE =
mengj@17998
    46
struct
mengj@17998
    47
blanchet@35865
    48
open Sledgehammer_FOL_Clause
blanchet@35865
    49
open Sledgehammer_Fact_Preprocessor
paulson@22825
    50
nipkow@31798
    51
(* Parameter t_full below indicates that full type information is to be
nipkow@31798
    52
exported *)
paulson@22825
    53
blanchet@35865
    54
(* If true, each function will be directly applied to as many arguments as
blanchet@35865
    55
   possible, avoiding use of the "apply" operator. Use of hBOOL is also
blanchet@35865
    56
   minimized. *)
wenzelm@30153
    57
val minimize_applies = true;
paulson@22064
    58
wenzelm@33035
    59
fun min_arity_of const_min_arity c = the_default 0 (Symtab.lookup const_min_arity c);
paulson@22064
    60
paulson@22825
    61
(*True if the constant ever appears outside of the top-level position in literals.
paulson@22825
    62
  If false, the constant always receives all of its arguments and is used as a predicate.*)
wenzelm@33035
    63
fun needs_hBOOL const_needs_hBOOL c =
wenzelm@33035
    64
  not minimize_applies orelse
wenzelm@33035
    65
    the_default false (Symtab.lookup const_needs_hBOOL c);
paulson@22064
    66
mengj@19198
    67
mengj@17998
    68
(******************************************************)
mengj@17998
    69
(* data types for typed combinator expressions        *)
mengj@17998
    70
(******************************************************)
mengj@17998
    71
mengj@17998
    72
type axiom_name = string;
mengj@17998
    73
type polarity = bool;
blanchet@35865
    74
type hol_clause_id = int;
mengj@17998
    75
blanchet@35865
    76
datatype combterm =
blanchet@35865
    77
  CombConst of string * fol_type * fol_type list (* Const and Free *) |
blanchet@35865
    78
  CombVar of string * fol_type |
blanchet@35865
    79
  CombApp of combterm * combterm
wenzelm@24311
    80
mengj@17998
    81
datatype literal = Literal of polarity * combterm;
mengj@17998
    82
blanchet@35865
    83
datatype hol_clause =
blanchet@35865
    84
  HOLClause of {clause_id: hol_clause_id, axiom_name: axiom_name, th: thm,
blanchet@35865
    85
                kind: kind, literals: literal list, ctypes_sorts: typ list};
mengj@17998
    86
mengj@17998
    87
mengj@17998
    88
(*********************************************************************)
mengj@17998
    89
(* convert a clause with type Term.term to a clause with type clause *)
mengj@17998
    90
(*********************************************************************)
mengj@17998
    91
paulson@21561
    92
fun isFalse (Literal(pol, CombConst(c,_,_))) =
blanchet@35865
    93
      (pol andalso c = "c_False") orelse (not pol andalso c = "c_True")
mengj@17998
    94
  | isFalse _ = false;
mengj@17998
    95
paulson@21561
    96
fun isTrue (Literal (pol, CombConst(c,_,_))) =
mengj@17998
    97
      (pol andalso c = "c_True") orelse
mengj@17998
    98
      (not pol andalso c = "c_False")
mengj@17998
    99
  | isTrue _ = false;
wenzelm@24311
   100
blanchet@35865
   101
fun isTaut (HOLClause {literals,...}) = exists isTrue literals;
mengj@17998
   102
immler@30151
   103
fun type_of dfg (Type (a, Ts)) =
immler@30151
   104
      let val (folTypes,ts) = types_of dfg Ts
blanchet@35865
   105
      in  (Comp(make_fixed_type_const dfg a, folTypes), ts)  end
blanchet@35865
   106
  | type_of _ (tp as TFree (a, _)) = (AtomF (make_fixed_type_var a), [tp])
blanchet@35865
   107
  | type_of _ (tp as TVar (v, _)) = (AtomV (make_schematic_type_var v), [tp])
immler@30151
   108
and types_of dfg Ts =
immler@30151
   109
      let val (folTyps,ts) = ListPair.unzip (map (type_of dfg) Ts)
blanchet@35865
   110
      in  (folTyps, union_all ts)  end;
mengj@17998
   111
mengj@17998
   112
(* same as above, but no gathering of sort information *)
immler@30151
   113
fun simp_type_of dfg (Type (a, Ts)) =
blanchet@35865
   114
      Comp(make_fixed_type_const dfg a, map (simp_type_of dfg) Ts)
blanchet@35865
   115
  | simp_type_of _ (TFree (a, _)) = AtomF (make_fixed_type_var a)
blanchet@35865
   116
  | simp_type_of _ (TVar (v, _)) = AtomV (make_schematic_type_var v);
mengj@18440
   117
mengj@18356
   118
immler@30151
   119
fun const_type_of dfg thy (c,t) =
immler@30151
   120
      let val (tp,ts) = type_of dfg t
immler@30151
   121
      in  (tp, ts, map (simp_type_of dfg) (Sign.const_typargs thy (c,t))) end;
mengj@18356
   122
mengj@17998
   123
(* convert a Term.term (with combinators) into a combterm, also accummulate sort info *)
immler@30151
   124
fun combterm_of dfg thy (Const(c,t)) =
immler@30151
   125
      let val (tp,ts,tvar_list) = const_type_of dfg thy (c,t)
blanchet@35865
   126
          val c' = CombConst(make_fixed_const dfg c, tp, tvar_list)
paulson@22078
   127
      in  (c',ts)  end
wenzelm@32994
   128
  | combterm_of dfg _ (Free(v,t)) =
immler@30151
   129
      let val (tp,ts) = type_of dfg t
blanchet@35865
   130
          val v' = CombConst(make_fixed_var v, tp, [])
paulson@22078
   131
      in  (v',ts)  end
wenzelm@32994
   132
  | combterm_of dfg _ (Var(v,t)) =
immler@30151
   133
      let val (tp,ts) = type_of dfg t
blanchet@35865
   134
          val v' = CombVar(make_schematic_var v,tp)
paulson@22078
   135
      in  (v',ts)  end
immler@30151
   136
  | combterm_of dfg thy (P $ Q) =
immler@30151
   137
      let val (P',tsP) = combterm_of dfg thy P
immler@30151
   138
          val (Q',tsQ) = combterm_of dfg thy Q
haftmann@33042
   139
      in  (CombApp(P',Q'), union (op =) tsP tsQ)  end
blanchet@35865
   140
  | combterm_of _ _ (t as Abs _) = raise CLAUSE ("HOL CLAUSE", t);
mengj@17998
   141
blanchet@35865
   142
fun predicate_of dfg thy ((@{const Not} $ P), polarity) = predicate_of dfg thy (P, not polarity)
immler@30151
   143
  | predicate_of dfg thy (t,polarity) = (combterm_of dfg thy (Envir.eta_contract t), polarity);
mengj@17998
   144
blanchet@35865
   145
fun literals_of_term1 dfg thy args (@{const Trueprop} $ P) = literals_of_term1 dfg thy args P
blanchet@35865
   146
  | literals_of_term1 dfg thy args (@{const "op |"} $ P $ Q) =
immler@30151
   147
      literals_of_term1 dfg thy (literals_of_term1 dfg thy args P) Q
immler@30151
   148
  | literals_of_term1 dfg thy (lits,ts) P =
immler@30151
   149
      let val ((pred,ts'),pol) = predicate_of dfg thy (P,true)
paulson@21509
   150
      in
haftmann@33042
   151
          (Literal(pol,pred)::lits, union (op =) ts ts')
paulson@21509
   152
      end;
mengj@17998
   153
immler@30151
   154
fun literals_of_term_dfg dfg thy P = literals_of_term1 dfg thy ([],[]) P;
immler@30151
   155
val literals_of_term = literals_of_term_dfg false;
mengj@17998
   156
blanchet@35865
   157
(* Trivial problem, which resolution cannot handle (empty clause) *)
blanchet@35865
   158
exception TRIVIAL;
wenzelm@28835
   159
mengj@17998
   160
(* making axiom and conjecture clauses *)
blanchet@35865
   161
fun make_clause dfg thy (clause_id, axiom_name, kind, th) =
immler@30151
   162
    let val (lits,ctypes_sorts) = literals_of_term_dfg dfg thy (prop_of th)
mengj@17998
   163
    in
blanchet@35865
   164
        if forall isFalse lits then
blanchet@35865
   165
            raise TRIVIAL
wenzelm@24311
   166
        else
blanchet@35865
   167
            HOLClause {clause_id = clause_id, axiom_name = axiom_name, th = th,
blanchet@35865
   168
                       kind = kind, literals = lits, ctypes_sorts = ctypes_sorts}
mengj@17998
   169
    end;
mengj@17998
   170
mengj@20016
   171
immler@30151
   172
fun add_axiom_clause dfg thy ((th,(name,id)), pairs) =
blanchet@35865
   173
  let val cls = make_clause dfg thy (id, name, Axiom, th)
paulson@21573
   174
  in
paulson@21573
   175
      if isTaut cls then pairs else (name,cls)::pairs
paulson@21573
   176
  end;
mengj@17998
   177
wenzelm@30193
   178
fun make_axiom_clauses dfg thy = List.foldl (add_axiom_clause dfg thy) [];
mengj@19354
   179
wenzelm@32994
   180
fun make_conjecture_clauses_aux _ _ _ [] = []
immler@30151
   181
  | make_conjecture_clauses_aux dfg thy n (th::ths) =
blanchet@35865
   182
      make_clause dfg thy (n,"conjecture", Conjecture, th) ::
immler@30151
   183
      make_conjecture_clauses_aux dfg thy (n+1) ths;
mengj@17998
   184
immler@30151
   185
fun make_conjecture_clauses dfg thy = make_conjecture_clauses_aux dfg thy 0;
mengj@17998
   186
mengj@17998
   187
mengj@17998
   188
(**********************************************************************)
mengj@17998
   189
(* convert clause into ATP specific formats:                          *)
mengj@17998
   190
(* TPTP used by Vampire and E                                         *)
mengj@19720
   191
(* DFG used by SPASS                                                  *)
mengj@17998
   192
(**********************************************************************)
mengj@17998
   193
paulson@22078
   194
(*Result of a function type; no need to check that the argument type matches.*)
blanchet@35865
   195
fun result_type (Comp ("tc_fun", [_, tp2])) = tp2
wenzelm@27187
   196
  | result_type _ = error "result_type"
paulson@22078
   197
wenzelm@32994
   198
fun type_of_combterm (CombConst (_, tp, _)) = tp
wenzelm@32994
   199
  | type_of_combterm (CombVar (_, tp)) = tp
wenzelm@32994
   200
  | type_of_combterm (CombApp (t1, _)) = result_type (type_of_combterm t1);
mengj@17998
   201
paulson@22064
   202
(*gets the head of a combinator application, along with the list of arguments*)
blanchet@35865
   203
fun strip_combterm_comb u =
paulson@22078
   204
    let fun stripc (CombApp(t,u), ts) = stripc (t, u::ts)
paulson@22064
   205
        |   stripc  x =  x
paulson@22064
   206
    in  stripc(u,[])  end;
paulson@22064
   207
paulson@22851
   208
val type_wrapper = "ti";
paulson@22851
   209
immler@30150
   210
fun head_needs_hBOOL const_needs_hBOOL (CombConst(c,_,_)) = needs_hBOOL const_needs_hBOOL c
wenzelm@32994
   211
  | head_needs_hBOOL _ _ = true;
paulson@22851
   212
nipkow@31791
   213
fun wrap_type t_full (s, tp) =
nipkow@31791
   214
  if t_full then
blanchet@35865
   215
      type_wrapper ^ paren_pack [s, string_of_fol_type tp]
paulson@22851
   216
  else s;
wenzelm@24311
   217
blanchet@35865
   218
fun apply ss = "hAPP" ^ paren_pack ss;
paulson@22851
   219
paulson@22064
   220
fun rev_apply (v, []) = v
paulson@22851
   221
  | rev_apply (v, arg::args) = apply [rev_apply (v, args), arg];
paulson@22064
   222
paulson@22064
   223
fun string_apply (v, args) = rev_apply (v, rev args);
paulson@22064
   224
paulson@22064
   225
(*Apply an operator to the argument strings, using either the "apply" operator or
paulson@22064
   226
  direct function application.*)
wenzelm@32994
   227
fun string_of_applic t_full cma (CombConst (c, _, tvars), args) =
paulson@22064
   228
      let val c = if c = "equal" then "c_fequal" else c
immler@30150
   229
          val nargs = min_arity_of cma c
paulson@22851
   230
          val args1 = List.take(args, nargs)
wenzelm@27187
   231
            handle Subscript => error ("string_of_applic: " ^ c ^ " has arity " ^
wenzelm@27187
   232
                                         Int.toString nargs ^ " but is applied to " ^
wenzelm@27187
   233
                                         space_implode ", " args)
paulson@22064
   234
          val args2 = List.drop(args, nargs)
blanchet@35865
   235
          val targs = if not t_full then map string_of_fol_type tvars else []
paulson@21513
   236
      in
blanchet@35865
   237
          string_apply (c ^ paren_pack (args1@targs), args2)
paulson@21513
   238
      end
wenzelm@32994
   239
  | string_of_applic _ _ (CombVar (v, _), args) = string_apply (v, args)
nipkow@31791
   240
  | string_of_applic _ _ _ = error "string_of_applic";
paulson@22064
   241
nipkow@31791
   242
fun wrap_type_if t_full cnh (head, s, tp) =
nipkow@31791
   243
  if head_needs_hBOOL cnh head then wrap_type t_full (s, tp) else s;
wenzelm@24311
   244
nipkow@31798
   245
fun string_of_combterm (params as (t_full, cma, cnh)) t =
blanchet@35865
   246
  let val (head, args) = strip_combterm_comb t
nipkow@31791
   247
  in  wrap_type_if t_full cnh (head,
nipkow@31798
   248
                    string_of_applic t_full cma (head, map (string_of_combterm (params)) args),
paulson@22851
   249
                    type_of_combterm t)
paulson@22851
   250
  end;
mengj@18356
   251
paulson@22064
   252
(*Boolean-valued terms are here converted to literals.*)
nipkow@31798
   253
fun boolify params t =
blanchet@35865
   254
  "hBOOL" ^ paren_pack [string_of_combterm params t];
paulson@22064
   255
nipkow@31798
   256
fun string_of_predicate (params as (_,_,cnh)) t =
paulson@22064
   257
  case t of
paulson@22078
   258
      (CombApp(CombApp(CombConst("equal",_,_), t1), t2)) =>
wenzelm@24311
   259
          (*DFG only: new TPTP prefers infix equality*)
blanchet@35865
   260
          ("equal" ^ paren_pack [string_of_combterm params t1, string_of_combterm params t2])
wenzelm@24311
   261
    | _ =>
blanchet@35865
   262
          case #1 (strip_combterm_comb t) of
nipkow@31798
   263
              CombConst(c,_,_) => if needs_hBOOL cnh c then boolify params t else string_of_combterm params t
nipkow@31798
   264
            | _ => boolify params t;
mengj@18356
   265
mengj@17998
   266
paulson@21561
   267
(*** tptp format ***)
mengj@19720
   268
nipkow@31798
   269
fun tptp_of_equality params pol (t1,t2) =
paulson@21513
   270
  let val eqop = if pol then " = " else " != "
nipkow@31798
   271
  in  string_of_combterm params t1 ^ eqop ^ string_of_combterm params t2  end;
paulson@21513
   272
blanchet@35865
   273
fun tptp_literal params (Literal(pol, CombApp(CombApp(CombConst("equal", _, _), t1), t2))) =
nipkow@31798
   274
      tptp_of_equality params pol (t1,t2)
nipkow@31798
   275
  | tptp_literal params (Literal(pol,pred)) =
blanchet@35865
   276
      tptp_sign pol (string_of_predicate params pred);
wenzelm@24311
   277
paulson@22064
   278
(*Given a clause, returns its literals paired with a list of literals concerning TFrees;
paulson@22064
   279
  the latter should only occur in conjecture clauses.*)
blanchet@35865
   280
fun tptp_type_lits params pos (HOLClause {literals, ctypes_sorts, ...}) =
nipkow@31798
   281
      (map (tptp_literal params) literals, 
blanchet@35865
   282
       map (tptp_of_typeLit pos) (add_typs ctypes_sorts));
wenzelm@24311
   283
blanchet@35865
   284
fun clause2tptp params (cls as HOLClause {axiom_name, clause_id, kind, ...}) =
blanchet@35865
   285
  let val (lits,tylits) = tptp_type_lits params (kind = Conjecture) cls
paulson@24937
   286
  in
blanchet@35865
   287
      (gen_tptp_cls (clause_id, axiom_name, kind, lits, tylits), tylits)
paulson@24937
   288
  end;
mengj@17998
   289
mengj@17998
   290
paulson@21561
   291
(*** dfg format ***)
paulson@21561
   292
blanchet@35865
   293
fun dfg_literal params (Literal(pol,pred)) = dfg_sign pol (string_of_predicate params pred);
mengj@19720
   294
blanchet@35865
   295
fun dfg_type_lits params pos (HOLClause {literals, ctypes_sorts, ...}) =
nipkow@31798
   296
      (map (dfg_literal params) literals, 
blanchet@35865
   297
       map (dfg_of_typeLit pos) (add_typs ctypes_sorts));
mengj@19720
   298
paulson@22078
   299
fun get_uvars (CombConst _) vars = vars
paulson@22078
   300
  | get_uvars (CombVar(v,_)) vars = (v::vars)
paulson@22078
   301
  | get_uvars (CombApp(P,Q)) vars = get_uvars P (get_uvars Q vars);
mengj@19720
   302
mengj@19720
   303
fun get_uvars_l (Literal(_,c)) = get_uvars c [];
mengj@19720
   304
blanchet@35865
   305
fun dfg_vars (HOLClause {literals,...}) = union_all (map get_uvars_l literals);
wenzelm@24311
   306
blanchet@35865
   307
fun clause2dfg params (cls as HOLClause {axiom_name, clause_id, kind,
blanchet@35865
   308
                                         ctypes_sorts, ...}) =
blanchet@35865
   309
  let val (lits,tylits) = dfg_type_lits params (kind = Conjecture) cls
paulson@24937
   310
      val vars = dfg_vars cls
blanchet@35865
   311
      val tvars = get_tvar_strs ctypes_sorts
paulson@24937
   312
  in
blanchet@35865
   313
      (gen_dfg_cls (clause_id, axiom_name, kind, lits, tylits, tvars@vars), tylits)
paulson@24937
   314
  end;
paulson@24937
   315
mengj@19720
   316
paulson@22064
   317
(** For DFG format: accumulate function and predicate declarations **)
mengj@19720
   318
blanchet@35865
   319
fun addtypes tvars tab = List.foldl add_foltype_funcs tab tvars;
mengj@19720
   320
wenzelm@32994
   321
fun add_decls (t_full, cma, cnh) (CombConst (c, _, tvars), (funcs, preds)) =
paulson@22064
   322
      if c = "equal" then (addtypes tvars funcs, preds)
paulson@21561
   323
      else
wenzelm@32962
   324
        let val arity = min_arity_of cma c
wenzelm@32962
   325
            val ntys = if not t_full then length tvars else 0
wenzelm@32962
   326
            val addit = Symtab.update(c, arity+ntys)
wenzelm@32962
   327
        in
wenzelm@32962
   328
            if needs_hBOOL cnh c then (addtypes tvars (addit funcs), preds)
wenzelm@32962
   329
            else (addtypes tvars funcs, addit preds)
wenzelm@32962
   330
        end
nipkow@31798
   331
  | add_decls _ (CombVar(_,ctp), (funcs,preds)) =
blanchet@35865
   332
      (add_foltype_funcs (ctp,funcs), preds)
nipkow@31798
   333
  | add_decls params (CombApp(P,Q),decls) = add_decls params (P,add_decls params (Q,decls));
mengj@19720
   334
blanchet@35865
   335
fun add_literal_decls params (Literal (_,c), decls) = add_decls params (c,decls);
mengj@19720
   336
blanchet@35865
   337
fun add_clause_decls params (HOLClause {literals, ...}, decls) =
nipkow@31798
   338
    List.foldl (add_literal_decls params) decls literals
wenzelm@27187
   339
    handle Symtab.DUP a => error ("function " ^ a ^ " has multiple arities")
mengj@19720
   340
nipkow@31798
   341
fun decls_of_clauses params clauses arity_clauses =
blanchet@35865
   342
  let val init_functab = Symtab.update (type_wrapper,2) (Symtab.update ("hAPP",2) init_functab)
paulson@22064
   343
      val init_predtab = Symtab.update ("hBOOL",1) Symtab.empty
nipkow@31798
   344
      val (functab,predtab) = (List.foldl (add_clause_decls params) (init_functab, init_predtab) clauses)
paulson@22064
   345
  in
blanchet@35865
   346
      (Symtab.dest (List.foldl add_arity_clause_funcs functab arity_clauses),
paulson@22064
   347
       Symtab.dest predtab)
paulson@22064
   348
  end;
mengj@19720
   349
blanchet@35865
   350
fun add_clause_preds (HOLClause {ctypes_sorts, ...}, preds) =
blanchet@35865
   351
  List.foldl add_type_sort_preds preds ctypes_sorts
wenzelm@27187
   352
  handle Symtab.DUP a => error ("predicate " ^ a ^ " has multiple arities")
paulson@21398
   353
paulson@21398
   354
(*Higher-order clauses have only the predicates hBOOL and type classes.*)
wenzelm@24311
   355
fun preds_of_clauses clauses clsrel_clauses arity_clauses =
mengj@19720
   356
    Symtab.dest
blanchet@35865
   357
        (List.foldl add_classrel_clause_preds
blanchet@35865
   358
               (List.foldl add_arity_clause_preds
wenzelm@30193
   359
                      (List.foldl add_clause_preds Symtab.empty clauses)
wenzelm@24311
   360
                      arity_clauses)
wenzelm@24311
   361
               clsrel_clauses)
mengj@19720
   362
mengj@18440
   363
mengj@18440
   364
(**********************************************************************)
mengj@19198
   365
(* write clauses to files                                             *)
mengj@19198
   366
(**********************************************************************)
mengj@19198
   367
paulson@21573
   368
val init_counters =
blanchet@35865
   369
  Symtab.make [("c_COMBI", 0), ("c_COMBK", 0), ("c_COMBB", 0), ("c_COMBC", 0),
blanchet@35865
   370
               ("c_COMBS", 0)];
wenzelm@24311
   371
wenzelm@32994
   372
fun count_combterm (CombConst (c, _, _), ct) =
paulson@21573
   373
     (case Symtab.lookup ct c of NONE => ct  (*no counter*)
paulson@21573
   374
                               | SOME n => Symtab.update (c,n+1) ct)
wenzelm@32994
   375
  | count_combterm (CombVar _, ct) = ct
paulson@22078
   376
  | count_combterm (CombApp(t1,t2), ct) = count_combterm(t1, count_combterm(t2, ct));
paulson@21573
   377
paulson@21573
   378
fun count_literal (Literal(_,t), ct) = count_combterm(t,ct);
paulson@21573
   379
blanchet@35865
   380
fun count_clause (HOLClause {literals, ...}, ct) =
blanchet@35865
   381
  List.foldl count_literal ct literals;
paulson@21573
   382
blanchet@35865
   383
fun count_user_clause user_lemmas (HOLClause {axiom_name, literals, ...}, ct) =
wenzelm@30193
   384
  if axiom_name mem_string user_lemmas then List.foldl count_literal ct literals
paulson@21573
   385
  else ct;
paulson@21573
   386
blanchet@35865
   387
fun cnf_helper_thms thy = cnf_rules_pairs thy o map pairname
mengj@20791
   388
immler@31752
   389
fun get_helper_clauses dfg thy isFO (conjectures, axcls, user_lemmas) =
blanchet@35865
   390
  if isFO then
blanchet@35865
   391
    []
paulson@23386
   392
  else
immler@31752
   393
    let
immler@31752
   394
        val axclauses = map #2 (make_axiom_clauses dfg thy axcls)
immler@31752
   395
        val ct0 = List.foldl count_clause init_counters conjectures
wenzelm@30193
   396
        val ct = List.foldl (count_user_clause user_lemmas) ct0 axclauses
wenzelm@33035
   397
        fun needed c = the (Symtab.lookup ct c) > 0
wenzelm@24311
   398
        val IK = if needed "c_COMBI" orelse needed "c_COMBK"
blanchet@35865
   399
                 then cnf_helper_thms thy [@{thm COMBI_def}, @{thm COMBK_def}]
paulson@21135
   400
                 else []
wenzelm@24311
   401
        val BC = if needed "c_COMBB" orelse needed "c_COMBC"
blanchet@35865
   402
                 then cnf_helper_thms thy [@{thm COMBB_def}, @{thm COMBC_def}]
wenzelm@24311
   403
                 else []
blanchet@35865
   404
        val S = if needed "c_COMBS" then cnf_helper_thms thy [@{thm COMBS_def}]
wenzelm@24311
   405
                else []
blanchet@35865
   406
        val other = cnf_helper_thms thy [@{thm fequal_imp_equal},
blanchet@35865
   407
                                         @{thm equal_imp_fequal}]
mengj@20791
   408
    in
immler@30151
   409
        map #2 (make_axiom_clauses dfg thy (other @ IK @ BC @ S))
paulson@23386
   410
    end;
mengj@20791
   411
paulson@22064
   412
(*Find the minimal arity of each function mentioned in the term. Also, note which uses
paulson@22064
   413
  are not at top level, to see if hBOOL is needed.*)
immler@30150
   414
fun count_constants_term toplev t (const_min_arity, const_needs_hBOOL) =
blanchet@35865
   415
  let val (head, args) = strip_combterm_comb t
paulson@22064
   416
      val n = length args
immler@30150
   417
      val (const_min_arity, const_needs_hBOOL) = fold (count_constants_term false) args (const_min_arity, const_needs_hBOOL)
paulson@22064
   418
  in
paulson@22064
   419
      case head of
wenzelm@24311
   420
          CombConst (a,_,_) => (*predicate or function version of "equal"?*)
wenzelm@24311
   421
            let val a = if a="equal" andalso not toplev then "c_fequal" else a
wenzelm@33029
   422
            val const_min_arity = Symtab.map_default (a, n) (Integer.min n) const_min_arity
wenzelm@24311
   423
            in
immler@30150
   424
              if toplev then (const_min_arity, const_needs_hBOOL)
immler@30150
   425
              else (const_min_arity, Symtab.update (a,true) (const_needs_hBOOL))
wenzelm@24311
   426
            end
wenzelm@32994
   427
        | _ => (const_min_arity, const_needs_hBOOL)
paulson@22064
   428
  end;
paulson@22064
   429
paulson@22064
   430
(*A literal is a top-level term*)
immler@30150
   431
fun count_constants_lit (Literal (_,t)) (const_min_arity, const_needs_hBOOL) =
immler@30150
   432
  count_constants_term true t (const_min_arity, const_needs_hBOOL);
paulson@22064
   433
blanchet@35865
   434
fun count_constants_clause (HOLClause {literals, ...})
blanchet@35865
   435
                           (const_min_arity, const_needs_hBOOL) =
immler@30150
   436
  fold count_constants_lit literals (const_min_arity, const_needs_hBOOL);
paulson@22064
   437
immler@30150
   438
fun display_arity const_needs_hBOOL (c,n) =
blanchet@35865
   439
  trace_msg (fn () => "Constant: " ^ c ^
blanchet@35826
   440
                " arity:\t" ^ Int.toString n ^
immler@30150
   441
                (if needs_hBOOL const_needs_hBOOL c then " needs hBOOL" else ""));
paulson@22064
   442
immler@31864
   443
fun count_constants (conjectures, _, extra_clauses, helper_clauses, _, _) =
wenzelm@30153
   444
  if minimize_applies then
immler@30150
   445
     let val (const_min_arity, const_needs_hBOOL) =
immler@30150
   446
          fold count_constants_clause conjectures (Symtab.empty, Symtab.empty)
immler@31864
   447
       |> fold count_constants_clause extra_clauses
immler@30149
   448
       |> fold count_constants_clause helper_clauses
immler@30150
   449
     val _ = List.app (display_arity const_needs_hBOOL) (Symtab.dest (const_min_arity))
immler@30150
   450
     in (const_min_arity, const_needs_hBOOL) end
immler@30150
   451
  else (Symtab.empty, Symtab.empty);
paulson@22064
   452
blanchet@35865
   453
(* TPTP format *)
immler@31749
   454
blanchet@35865
   455
fun write_tptp_file t_full file clauses =
immler@31409
   456
  let
immler@31864
   457
    val (conjectures, axclauses, _, helper_clauses,
immler@31864
   458
      classrel_clauses, arity_clauses) = clauses
nipkow@31791
   459
    val (cma, cnh) = count_constants clauses
nipkow@31791
   460
    val params = (t_full, cma, cnh)
nipkow@31791
   461
    val (tptp_clss,tfree_litss) = ListPair.unzip (map (clause2tptp params) conjectures)
blanchet@35865
   462
    val tfree_clss = map tptp_tfree_clause (List.foldl (uncurry (union (op =))) [] tfree_litss)
immler@31832
   463
    val _ =
immler@31832
   464
      File.write_list file (
immler@31832
   465
        map (#1 o (clause2tptp params)) axclauses @
immler@31832
   466
        tfree_clss @
immler@31832
   467
        tptp_clss @
blanchet@35865
   468
        map tptp_classrel_clause classrel_clauses @
blanchet@35865
   469
        map tptp_arity_clause arity_clauses @
immler@31832
   470
        map (#1 o (clause2tptp params)) helper_clauses)
immler@31832
   471
    in (length axclauses + 1, length tfree_clss + length tptp_clss)
immler@31409
   472
  end;
mengj@19198
   473
mengj@19720
   474
blanchet@35865
   475
(* DFG format *)
mengj@19720
   476
blanchet@35865
   477
fun write_dfg_file t_full file clauses =
immler@31409
   478
  let
immler@31864
   479
    val (conjectures, axclauses, _, helper_clauses,
immler@31864
   480
      classrel_clauses, arity_clauses) = clauses
nipkow@31791
   481
    val (cma, cnh) = count_constants clauses
nipkow@31791
   482
    val params = (t_full, cma, cnh)
nipkow@31791
   483
    val (dfg_clss, tfree_litss) = ListPair.unzip (map (clause2dfg params) conjectures)
immler@31831
   484
    and probname = Path.implode (Path.base file)
nipkow@31791
   485
    val axstrs = map (#1 o (clause2dfg params)) axclauses
blanchet@35865
   486
    val tfree_clss = map dfg_tfree_clause (union_all tfree_litss)
nipkow@31791
   487
    val helper_clauses_strs = map (#1 o (clause2dfg params)) helper_clauses
nipkow@31791
   488
    val (funcs,cl_preds) = decls_of_clauses params (helper_clauses @ conjectures @ axclauses) arity_clauses
immler@31409
   489
    and ty_preds = preds_of_clauses axclauses classrel_clauses arity_clauses
immler@31832
   490
    val _ =
immler@31832
   491
      File.write_list file (
blanchet@35865
   492
        string_of_start probname ::
blanchet@35865
   493
        string_of_descrip probname ::
blanchet@35865
   494
        string_of_symbols (string_of_funcs funcs)
blanchet@35865
   495
          (string_of_preds (cl_preds @ ty_preds)) ::
immler@31832
   496
        "list_of_clauses(axioms,cnf).\n" ::
immler@31832
   497
        axstrs @
blanchet@35865
   498
        map dfg_classrel_clause classrel_clauses @
blanchet@35865
   499
        map dfg_arity_clause arity_clauses @
immler@31832
   500
        helper_clauses_strs @
immler@31832
   501
        ["end_of_list.\n\nlist_of_clauses(conjectures,cnf).\n"] @
immler@31832
   502
        tfree_clss @
immler@31832
   503
        dfg_clss @
immler@31832
   504
        ["end_of_list.\n\n",
immler@31832
   505
        (*VarWeight=3 helps the HO problems, probably by counteracting the presence of hAPP*)
immler@31832
   506
         "list_of_settings(SPASS).\n{*\nset_flag(VarWeight,3).\n*}\nend_of_list.\n\n",
immler@31832
   507
         "end_problem.\n"])
immler@31832
   508
immler@31832
   509
    in (length axclauses + length classrel_clauses + length arity_clauses +
immler@31832
   510
      length helper_clauses + 1, length tfree_clss + length dfg_clss)
immler@31409
   511
  end;
mengj@19720
   512
wenzelm@33319
   513
end;