src/HOL/Tools/ATP/atp_problem.ML
author blanchet
Sun, 22 May 2011 14:51:42 +0200
changeset 43785 9e620869a576
parent 43783 ad34216cff2f
child 43802 f30ae82cb62e
permissions -rw-r--r--
improved Waldmeister support -- even run it by default on unit equational goals
     1 (*  Title:      HOL/Tools/ATP/atp_problem.ML
     2     Author:     Jia Meng, Cambridge University Computer Laboratory and NICTA
     3     Author:     Jasmin Blanchette, TU Muenchen
     4 
     5 Abstract representation of ATP problems and TPTP syntax.
     6 *)
     7 
     8 signature ATP_PROBLEM =
     9 sig
    10   datatype 'a fo_term = ATerm of 'a * 'a fo_term list
    11   datatype quantifier = AForall | AExists
    12   datatype connective = ANot | AAnd | AOr | AImplies | AIf | AIff | ANotIff
    13   datatype ('a, 'b, 'c) formula =
    14     AQuant of quantifier * ('a * 'b option) list * ('a, 'b, 'c) formula |
    15     AConn of connective * ('a, 'b, 'c) formula list |
    16     AAtom of 'c
    17 
    18   datatype format = CNF_UEQ | FOF | TFF
    19   datatype formula_kind = Axiom | Definition | Lemma | Hypothesis | Conjecture
    20   datatype 'a problem_line =
    21     Decl of string * 'a * 'a list * 'a |
    22     Formula of string * formula_kind * ('a, 'a, 'a fo_term) formula
    23                * string fo_term option * string fo_term option
    24   type 'a problem = (string * 'a problem_line list) list
    25 
    26   (* official TPTP syntax *)
    27   val tptp_special_prefix : string
    28   val tptp_false : string
    29   val tptp_true : string
    30   val tptp_tff_type_of_types : string
    31   val tptp_tff_bool_type : string
    32   val tptp_tff_individual_type : string
    33   val is_atp_variable : string -> bool
    34   val mk_anot : ('a, 'b, 'c) formula -> ('a, 'b, 'c) formula
    35   val mk_aconn :
    36     connective -> ('a, 'b, 'c) formula -> ('a, 'b, 'c) formula
    37     -> ('a, 'b, 'c) formula
    38   val formula_map : ('c -> 'd) -> ('a, 'b, 'c) formula -> ('a, 'b, 'd) formula
    39   val timestamp : unit -> string
    40   val hashw : word * word -> word
    41   val hashw_string : string * word -> word
    42   val tptp_strings_for_atp_problem : format -> string problem -> string list
    43   val filter_cnf_ueq_problem :
    44     (string * string) problem -> (string * string) problem
    45   val nice_atp_problem :
    46     bool -> ('a * (string * string) problem_line list) list
    47     -> ('a * string problem_line list) list
    48        * (string Symtab.table * string Symtab.table) option
    49 end;
    50 
    51 structure ATP_Problem : ATP_PROBLEM =
    52 struct
    53 
    54 (** ATP problem **)
    55 
    56 datatype 'a fo_term = ATerm of 'a * 'a fo_term list
    57 datatype quantifier = AForall | AExists
    58 datatype connective = ANot | AAnd | AOr | AImplies | AIf | AIff | ANotIff
    59 datatype ('a, 'b, 'c) formula =
    60   AQuant of quantifier * ('a * 'b option) list * ('a, 'b, 'c) formula |
    61   AConn of connective * ('a, 'b, 'c) formula list |
    62   AAtom of 'c
    63 
    64 datatype format = CNF_UEQ | FOF | TFF
    65 datatype formula_kind = Axiom | Definition | Lemma | Hypothesis | Conjecture
    66 datatype 'a problem_line =
    67   Decl of string * 'a * 'a list * 'a |
    68   Formula of string * formula_kind * ('a, 'a, 'a fo_term) formula
    69              * string fo_term option * string fo_term option
    70 type 'a problem = (string * 'a problem_line list) list
    71 
    72 (* official TPTP syntax *)
    73 val tptp_special_prefix = "$"
    74 val tptp_false = "$false"
    75 val tptp_true = "$true"
    76 val tptp_tff_type_of_types = "$tType"
    77 val tptp_tff_bool_type = "$o"
    78 val tptp_tff_individual_type = "$i"
    79 
    80 fun is_atp_variable s = Char.isUpper (String.sub (s, 0))
    81 
    82 fun mk_anot (AConn (ANot, [phi])) = phi
    83   | mk_anot phi = AConn (ANot, [phi])
    84 fun mk_aconn c phi1 phi2 = AConn (c, [phi1, phi2])
    85 
    86 fun formula_map f (AQuant (q, xs, phi)) = AQuant (q, xs, formula_map f phi)
    87   | formula_map f (AConn (c, phis)) = AConn (c, map (formula_map f) phis)
    88   | formula_map f (AAtom tm) = AAtom (f tm)
    89 
    90 val timestamp = Date.fmt "%Y-%m-%d %H:%M:%S" o Date.fromTimeLocal o Time.now
    91 
    92 (* This hash function is recommended in Compilers: Principles, Techniques, and
    93    Tools, by Aho, Sethi, and Ullman. The "hashpjw" function, which they
    94    particularly recommend, triggers a bug in versions of Poly/ML up to 4.2.0. *)
    95 fun hashw (u, w) = Word.+ (u, Word.* (0w65599, w))
    96 fun hashw_char (c, w) = hashw (Word.fromInt (Char.ord c), w)
    97 fun hashw_string (s : string, w) = CharVector.foldl hashw_char w s
    98 
    99 fun string_for_kind Axiom = "axiom"
   100   | string_for_kind Definition = "definition"
   101   | string_for_kind Lemma = "lemma"
   102   | string_for_kind Hypothesis = "hypothesis"
   103   | string_for_kind Conjecture = "conjecture"
   104 
   105 fun string_for_term (ATerm (s, [])) = s
   106   | string_for_term (ATerm ("equal", ts)) =
   107     space_implode " = " (map string_for_term ts)
   108   | string_for_term (ATerm ("[]", ts)) =
   109     (* used for lists in the optional "source" field of a derivation *)
   110     "[" ^ commas (map string_for_term ts) ^ "]"
   111   | string_for_term (ATerm (s, ts)) =
   112     s ^ "(" ^ commas (map string_for_term ts) ^ ")"
   113 fun string_for_quantifier AForall = "!"
   114   | string_for_quantifier AExists = "?"
   115 fun string_for_connective ANot = "~"
   116   | string_for_connective AAnd = "&"
   117   | string_for_connective AOr = "|"
   118   | string_for_connective AImplies = "=>"
   119   | string_for_connective AIf = "<="
   120   | string_for_connective AIff = "<=>"
   121   | string_for_connective ANotIff = "<~>"
   122 fun string_for_bound_var TFF (s, ty) =
   123     s ^ " : " ^ (ty |> the_default tptp_tff_individual_type)
   124   | string_for_bound_var _ (s, _) = s
   125 fun string_for_formula format (AQuant (q, xs, phi)) =
   126     "(" ^ string_for_quantifier q ^
   127     "[" ^ commas (map (string_for_bound_var format) xs) ^ "] : " ^
   128     string_for_formula format phi ^ ")"
   129   | string_for_formula _ (AConn (ANot, [AAtom (ATerm ("equal", ts))])) =
   130     space_implode " != " (map string_for_term ts)
   131   | string_for_formula format (AConn (c, [phi])) =
   132     "(" ^ string_for_connective c ^ " " ^ string_for_formula format phi ^ ")"
   133   | string_for_formula format (AConn (c, phis)) =
   134     "(" ^ space_implode (" " ^ string_for_connective c ^ " ")
   135                         (map (string_for_formula format) phis) ^ ")"
   136   | string_for_formula _ (AAtom tm) = string_for_term tm
   137 
   138 fun string_for_symbol_type [] res_ty = res_ty
   139   | string_for_symbol_type [arg_ty] res_ty = arg_ty ^ " > " ^ res_ty
   140   | string_for_symbol_type arg_tys res_ty =
   141     string_for_symbol_type ["(" ^ space_implode " * " arg_tys ^ ")"] res_ty
   142 
   143 val default_source =
   144   ATerm ("inference", ATerm ("isabelle", []) :: replicate 2 (ATerm ("[]", [])))
   145 
   146 fun string_for_problem_line _ (Decl (ident, sym, arg_tys, res_ty)) =
   147     "tff(" ^ ident ^ ", type,\n    " ^ sym ^ " : " ^
   148     string_for_symbol_type arg_tys res_ty ^ ").\n"
   149   | string_for_problem_line format (Formula (ident, kind, phi, source, info)) =
   150     (case format of CNF_UEQ => "cnf" | FOF => "fof" | TFF => "tff") ^
   151     "(" ^ ident ^ ", " ^ string_for_kind kind ^ ",\n    (" ^
   152     string_for_formula format phi ^ ")" ^
   153     (case (source, info) of
   154        (NONE, NONE) => ""
   155      | (SOME tm, NONE) => ", " ^ string_for_term tm
   156      | (_, SOME tm) =>
   157        ", " ^ string_for_term (source |> the_default default_source) ^
   158        ", " ^ string_for_term tm) ^ ").\n"
   159 fun tptp_strings_for_atp_problem format problem =
   160   "% This file was generated by Isabelle (most likely Sledgehammer)\n\
   161   \% " ^ timestamp () ^ "\n" ::
   162   maps (fn (_, []) => []
   163          | (heading, lines) =>
   164            "\n% " ^ heading ^ " (" ^ string_of_int (length lines) ^ ")\n" ::
   165            map (string_for_problem_line format) lines)
   166        problem
   167 
   168 
   169 (** CNF UEQ (Waldmeister) **)
   170 
   171 exception LOST_CONJECTURE of unit
   172 
   173 fun is_problem_line_negated (Formula (_, _, AConn (ANot, _), _, _)) = true
   174   | is_problem_line_negated _ = false
   175 
   176 fun is_problem_line_cnf_ueq
   177         (Formula (_, _, AAtom (ATerm (("equal", _), _)), _, _)) = true
   178   | is_problem_line_cnf_ueq _ = false
   179 
   180 fun open_conjecture_term (ATerm ((s, s'), tms)) =
   181   ATerm (s |> is_atp_variable s ? Name.desymbolize false |> `I,
   182          tms |> map open_conjecture_term)
   183 fun open_formula conj (AQuant (AForall, _, phi)) = open_formula conj phi
   184   | open_formula true (AAtom t) = AAtom (open_conjecture_term t)
   185   | open_formula _ phi = phi
   186 fun open_formula_line (Formula (ident, kind, phi, source, info)) =
   187     Formula (ident, kind, open_formula (kind = Conjecture) phi, source, info)
   188   | open_formula_line line = line
   189 
   190 fun negate_conjecture_line (Formula (ident, Conjecture, phi, source, info)) =
   191     Formula (ident, Hypothesis, mk_anot phi, source, info)
   192   | negate_conjecture_line line = line
   193 
   194 val filter_cnf_ueq_problem =
   195   map (apsnd (map open_formula_line
   196               #> filter is_problem_line_cnf_ueq
   197               #> map negate_conjecture_line))
   198   #> (fn problem =>
   199          let
   200            val conjs = problem |> maps snd |> filter is_problem_line_negated
   201          in if length conjs = 1 then problem else [] end)
   202 
   203 
   204 (** Nice names **)
   205 
   206 fun empty_name_pool readable_names =
   207   if readable_names then SOME (Symtab.empty, Symtab.empty) else NONE
   208 
   209 fun pool_fold f xs z = pair z #> fold_rev (fn x => uncurry (f x)) xs
   210 fun pool_map f xs =
   211   pool_fold (fn x => fn ys => fn pool => f x pool |>> (fn y => y :: ys)) xs []
   212 
   213 val no_qualifiers =
   214   let
   215     fun skip [] = []
   216       | skip (#"." :: cs) = skip cs
   217       | skip (c :: cs) = if Char.isAlphaNum c then skip cs else c :: keep cs
   218     and keep [] = []
   219       | keep (#"." :: cs) = skip cs
   220       | keep (c :: cs) = c :: keep cs
   221   in String.explode #> rev #> keep #> rev #> String.implode end
   222 
   223 (* Long names can slow down the ATPs. *)
   224 val max_readable_name_size = 20
   225 
   226 (* "op" is also reserved, to avoid the unreadable "op_1", "op_2", etc., in the
   227    problem files. "equal" is reserved by some ATPs. "eq" is reserved to ensure
   228    that "HOL.eq" is correctly mapped to equality. *)
   229 val reserved_nice_names = ["op", "equal", "eq"]
   230 
   231 fun readable_name full_name s =
   232   if s = full_name then
   233     s
   234   else
   235     s |> no_qualifiers
   236       |> Name.desymbolize (Char.isUpper (String.sub (full_name, 0)))
   237       |> (fn s =>
   238              if size s > max_readable_name_size then
   239                String.substring (s, 0, max_readable_name_size div 2 - 4) ^
   240                Word.toString (hashw_string (full_name, 0w0)) ^
   241                String.extract (s, size s - max_readable_name_size div 2 + 4,
   242                                NONE)
   243              else
   244                s)
   245       |> (fn s => if member (op =) reserved_nice_names s then full_name else s)
   246 
   247 fun nice_name (full_name, _) NONE = (full_name, NONE)
   248   | nice_name (full_name, desired_name) (SOME the_pool) =
   249     if String.isPrefix "$" full_name then
   250       (full_name, SOME the_pool)
   251     else case Symtab.lookup (fst the_pool) full_name of
   252       SOME nice_name => (nice_name, SOME the_pool)
   253     | NONE =>
   254       let
   255         val nice_prefix = readable_name full_name desired_name
   256         fun add j =
   257           let
   258             val nice_name =
   259               nice_prefix ^ (if j = 0 then "" else "_" ^ string_of_int j)
   260           in
   261             case Symtab.lookup (snd the_pool) nice_name of
   262               SOME full_name' =>
   263               if full_name = full_name' then (nice_name, the_pool)
   264               else add (j + 1)
   265             | NONE =>
   266               (nice_name,
   267                (Symtab.update_new (full_name, nice_name) (fst the_pool),
   268                 Symtab.update_new (nice_name, full_name) (snd the_pool)))
   269           end
   270       in add 0 |> apsnd SOME end
   271 
   272 fun nice_term (ATerm (name, ts)) =
   273   nice_name name ##>> pool_map nice_term ts #>> ATerm
   274 fun nice_formula (AQuant (q, xs, phi)) =
   275     pool_map nice_name (map fst xs)
   276     ##>> pool_map (fn NONE => pair NONE
   277                     | SOME ty => nice_name ty #>> SOME) (map snd xs)
   278     ##>> nice_formula phi
   279     #>> (fn ((ss, ts), phi) => AQuant (q, ss ~~ ts, phi))
   280   | nice_formula (AConn (c, phis)) =
   281     pool_map nice_formula phis #>> curry AConn c
   282   | nice_formula (AAtom tm) = nice_term tm #>> AAtom
   283 fun nice_problem_line (Decl (ident, sym, arg_tys, res_ty)) =
   284     nice_name sym
   285     ##>> pool_map nice_name arg_tys
   286     ##>> nice_name res_ty
   287     #>> (fn ((sym, arg_tys), res_ty) => Decl (ident, sym, arg_tys, res_ty))
   288   | nice_problem_line (Formula (ident, kind, phi, source, info)) =
   289     nice_formula phi #>> (fn phi => Formula (ident, kind, phi, source, info))
   290 fun nice_problem problem =
   291   pool_map (fn (heading, lines) =>
   292                pool_map nice_problem_line lines #>> pair heading) problem
   293 fun nice_atp_problem readable_names problem =
   294   nice_problem problem (empty_name_pool readable_names)
   295 
   296 end;