src/HOL/Tools/Sledgehammer/sledgehammer_tptp_format.ML
changeset 38018 ee6c11ae8077
parent 38017 3ad3e3ca2451
child 38019 e207a64e1e0b
equal deleted inserted replaced
38017:3ad3e3ca2451 38018:ee6c11ae8077
     1 (*  Title:      HOL/Tools/Sledgehammer/sledgehammer_tptp_format.ML
       
     2     Author:     Jia Meng, NICTA
       
     3     Author:     Jasmin Blanchette, TU Muenchen
       
     4 
       
     5 TPTP syntax.
       
     6 *)
       
     7 
       
     8 signature SLEDGEHAMMER_TPTP_FORMAT =
       
     9 sig
       
    10   type kind = Metis_Clauses.kind
       
    11 
       
    12   datatype 'a fo_term = ATerm of 'a * 'a fo_term list
       
    13   datatype quantifier = AForall | AExists
       
    14   datatype connective = ANot | AAnd | AOr | AImplies | AIf | AIff | ANotIff
       
    15   datatype ('a, 'b) formula =
       
    16     AQuant of quantifier * 'a list * ('a, 'b) formula |
       
    17     AConn of connective * ('a, 'b) formula list |
       
    18     APred of 'b
       
    19 
       
    20   datatype 'a problem_line = Fof of string * kind * ('a, 'a fo_term) formula
       
    21   type 'a problem = (string * 'a problem_line list) list
       
    22 
       
    23   val is_tptp_variable : string -> bool
       
    24   val strings_for_tptp_problem :
       
    25     (string * string problem_line list) list -> string list
       
    26   val nice_tptp_problem :
       
    27     bool -> ('a * (string * string) problem_line list) list
       
    28     -> ('a * string problem_line list) list
       
    29        * (string Symtab.table * string Symtab.table) option
       
    30 end;
       
    31 
       
    32 structure Sledgehammer_TPTP_Format : SLEDGEHAMMER_TPTP_FORMAT =
       
    33 struct
       
    34 
       
    35 open Sledgehammer_Util
       
    36 
       
    37 type kind = Metis_Clauses.kind
       
    38 
       
    39 (** ATP problem **)
       
    40 
       
    41 datatype 'a fo_term = ATerm of 'a * 'a fo_term list
       
    42 datatype quantifier = AForall | AExists
       
    43 datatype connective = ANot | AAnd | AOr | AImplies | AIf | AIff | ANotIff
       
    44 datatype ('a, 'b) formula =
       
    45   AQuant of quantifier * 'a list * ('a, 'b) formula |
       
    46   AConn of connective * ('a, 'b) formula list |
       
    47   APred of 'b
       
    48 
       
    49 datatype 'a problem_line = Fof of string * kind * ('a, 'a fo_term) formula
       
    50 type 'a problem = (string * 'a problem_line list) list
       
    51 
       
    52 fun string_for_term (ATerm (s, [])) = s
       
    53   | string_for_term (ATerm (s, ts)) =
       
    54     if s = "equal" then space_implode " = " (map string_for_term ts)
       
    55     else s ^ "(" ^ commas (map string_for_term ts) ^ ")"
       
    56 fun string_for_quantifier AForall = "!"
       
    57   | string_for_quantifier AExists = "?"
       
    58 fun string_for_connective ANot = "~"
       
    59   | string_for_connective AAnd = "&"
       
    60   | string_for_connective AOr = "|"
       
    61   | string_for_connective AImplies = "=>"
       
    62   | string_for_connective AIf = "<="
       
    63   | string_for_connective AIff = "<=>"
       
    64   | string_for_connective ANotIff = "<~>"
       
    65 fun string_for_formula (AQuant (q, xs, phi)) =
       
    66     string_for_quantifier q ^ "[" ^ commas xs ^ "]: " ^
       
    67     string_for_formula phi
       
    68   | string_for_formula (AConn (ANot, [APred (ATerm ("equal", ts))])) =
       
    69     space_implode " != " (map string_for_term ts)
       
    70   | string_for_formula (AConn (c, [phi])) =
       
    71     string_for_connective c ^ " " ^ string_for_formula phi
       
    72   | string_for_formula (AConn (c, phis)) =
       
    73     "(" ^ space_implode (" " ^ string_for_connective c ^ " ")
       
    74                         (map string_for_formula phis) ^ ")"
       
    75   | string_for_formula (APred tm) = string_for_term tm
       
    76 
       
    77 fun string_for_problem_line (Fof (ident, kind, phi)) =
       
    78   "fof(" ^ ident ^ ", " ^
       
    79   (case kind of Axiom => "axiom" | Conjecture => "conjecture") ^ ",\n" ^
       
    80   "    (" ^ string_for_formula phi ^ ")).\n"
       
    81 fun strings_for_tptp_problem problem =
       
    82   "% This file was generated by Isabelle (most likely Sledgehammer)\n\
       
    83   \% " ^ timestamp () ^ "\n" ::
       
    84   maps (fn (_, []) => []
       
    85          | (heading, lines) =>
       
    86            "\n% " ^ heading ^ " (" ^ Int.toString (length lines) ^ ")\n" ::
       
    87            map string_for_problem_line lines) problem
       
    88 
       
    89 fun is_tptp_variable s = Char.isUpper (String.sub (s, 0))
       
    90 
       
    91 
       
    92 (** Nice names **)
       
    93 
       
    94 fun empty_name_pool readable_names =
       
    95   if readable_names then SOME (Symtab.empty, Symtab.empty) else NONE
       
    96 
       
    97 fun pool_fold f xs z = pair z #> fold_rev (fn x => uncurry (f x)) xs
       
    98 fun pool_map f xs =
       
    99   pool_fold (fn x => fn ys => fn pool => f x pool |>> (fn y => y :: ys)) xs []
       
   100 
       
   101 (* "equal" is reserved by some ATPs. "op" is also reserved, to avoid the
       
   102    unreadable "op_1", "op_2", etc., in the problem files. *)
       
   103 val reserved_nice_names = ["equal", "op"]
       
   104 fun readable_name full_name s =
       
   105   if s = full_name then
       
   106     s
       
   107   else
       
   108     let
       
   109       val s = s |> Long_Name.base_name
       
   110                 |> Name.desymbolize (Char.isUpper (String.sub (full_name, 0)))
       
   111     in if member (op =) reserved_nice_names s then full_name else s end
       
   112 
       
   113 fun nice_name (full_name, _) NONE = (full_name, NONE)
       
   114   | nice_name (full_name, desired_name) (SOME the_pool) =
       
   115     case Symtab.lookup (fst the_pool) full_name of
       
   116       SOME nice_name => (nice_name, SOME the_pool)
       
   117     | NONE =>
       
   118       let
       
   119         val nice_prefix = readable_name full_name desired_name
       
   120         fun add j =
       
   121           let
       
   122             val nice_name = nice_prefix ^
       
   123                             (if j = 0 then "" else "_" ^ Int.toString j)
       
   124           in
       
   125             case Symtab.lookup (snd the_pool) nice_name of
       
   126               SOME full_name' =>
       
   127               if full_name = full_name' then (nice_name, the_pool)
       
   128               else add (j + 1)
       
   129             | NONE =>
       
   130               (nice_name,
       
   131                (Symtab.update_new (full_name, nice_name) (fst the_pool),
       
   132                 Symtab.update_new (nice_name, full_name) (snd the_pool)))
       
   133           end
       
   134       in add 0 |> apsnd SOME end
       
   135 
       
   136 
       
   137 fun nice_term (ATerm (name, ts)) =
       
   138   nice_name name ##>> pool_map nice_term ts #>> ATerm
       
   139 fun nice_formula (AQuant (q, xs, phi)) =
       
   140     pool_map nice_name xs ##>> nice_formula phi
       
   141     #>> (fn (xs, phi) => AQuant (q, xs, phi))
       
   142   | nice_formula (AConn (c, phis)) =
       
   143     pool_map nice_formula phis #>> curry AConn c
       
   144   | nice_formula (APred tm) = nice_term tm #>> APred
       
   145 fun nice_problem_line (Fof (ident, kind, phi)) =
       
   146   nice_formula phi #>> (fn phi => Fof (ident, kind, phi))
       
   147 fun nice_problem problem =
       
   148   pool_map (fn (heading, lines) =>
       
   149                pool_map nice_problem_line lines #>> pair heading) problem
       
   150 fun nice_tptp_problem readable_names problem =
       
   151   nice_problem problem (empty_name_pool readable_names)
       
   152 
       
   153 end;