src/HOL/Tools/ATP/atp_problem.ML
author blanchet
Thu, 21 Apr 2011 22:18:28 +0200
changeset 43320 494e4ac5b0f8
parent 43088 662b50b7126f
child 43396 7a506b0b644f
permissions -rw-r--r--
detect some unsound proofs before showing them to the user
     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) formula =
    14     AQuant of quantifier * 'a list * ('a, 'b) formula |
    15     AConn of connective * ('a, 'b) formula list |
    16     AAtom of 'b
    17   type 'a uniform_formula = ('a, 'a fo_term) formula
    18 
    19   datatype kind = Axiom | Definition | Lemma | Hypothesis | Conjecture
    20   datatype 'a problem_line =
    21     Fof of string * kind * ('a, 'a fo_term) formula * string fo_term option
    22   type 'a problem = (string * 'a problem_line list) list
    23 
    24   val timestamp : unit -> string
    25   val is_atp_variable : string -> bool
    26   val tptp_strings_for_atp_problem :
    27     bool -> (string * string problem_line list) list -> string list
    28   val nice_atp_problem :
    29     bool -> ('a * (string * string) problem_line list) list
    30     -> ('a * string problem_line list) list
    31        * (string Symtab.table * string Symtab.table) option
    32 end;
    33 
    34 structure ATP_Problem : ATP_PROBLEM =
    35 struct
    36 
    37 (** ATP problem **)
    38 
    39 datatype 'a fo_term = ATerm of 'a * 'a fo_term list
    40 datatype quantifier = AForall | AExists
    41 datatype connective = ANot | AAnd | AOr | AImplies | AIf | AIff | ANotIff
    42 datatype ('a, 'b) formula =
    43   AQuant of quantifier * 'a list * ('a, 'b) formula |
    44   AConn of connective * ('a, 'b) formula list |
    45   AAtom of 'b
    46 type 'a uniform_formula = ('a, 'a fo_term) formula
    47 
    48 datatype kind = Axiom | Definition | Lemma | Hypothesis | Conjecture
    49 datatype 'a problem_line =
    50   Fof of string * kind * ('a, 'a fo_term) formula * string fo_term option
    51 type 'a problem = (string * 'a problem_line list) list
    52 
    53 val timestamp = Date.fmt "%Y-%m-%d %H:%M:%S" o Date.fromTimeLocal o Time.now
    54 
    55 fun string_for_kind Axiom = "axiom"
    56   | string_for_kind Definition = "definition"
    57   | string_for_kind Lemma = "lemma"
    58   | string_for_kind Hypothesis = "hypothesis"
    59   | string_for_kind Conjecture = "conjecture"
    60 
    61 fun string_for_term (ATerm (s, [])) = s
    62   | string_for_term (ATerm ("equal", ts)) =
    63     space_implode " = " (map string_for_term ts)
    64   | string_for_term (ATerm ("[]", ts)) =
    65     (* used for lists in the optional "source" field of a derivation *)
    66     "[" ^ commas (map string_for_term ts) ^ "]"
    67   | string_for_term (ATerm (s, ts)) =
    68     s ^ "(" ^ commas (map string_for_term ts) ^ ")"
    69 fun string_for_quantifier AForall = "!"
    70   | string_for_quantifier AExists = "?"
    71 fun string_for_connective ANot = "~"
    72   | string_for_connective AAnd = "&"
    73   | string_for_connective AOr = "|"
    74   | string_for_connective AImplies = "=>"
    75   | string_for_connective AIf = "<="
    76   | string_for_connective AIff = "<=>"
    77   | string_for_connective ANotIff = "<~>"
    78 fun string_for_formula (AQuant (q, xs, phi)) =
    79     "(" ^ string_for_quantifier q ^ "[" ^ commas xs ^ "] : " ^
    80     string_for_formula phi ^ ")"
    81   | string_for_formula (AConn (ANot, [AAtom (ATerm ("equal", ts))])) =
    82     space_implode " != " (map string_for_term ts)
    83   | string_for_formula (AConn (c, [phi])) =
    84     "(" ^ string_for_connective c ^ " " ^ string_for_formula phi ^ ")"
    85   | string_for_formula (AConn (c, phis)) =
    86     "(" ^ space_implode (" " ^ string_for_connective c ^ " ")
    87                         (map string_for_formula phis) ^ ")"
    88   | string_for_formula (AAtom tm) = string_for_term tm
    89 
    90 fun string_for_problem_line use_conjecture_for_hypotheses
    91                             (Fof (ident, kind, phi, source)) =
    92   let
    93     val (kind, phi) =
    94       if kind = Hypothesis andalso use_conjecture_for_hypotheses then
    95         (Conjecture, AConn (ANot, [phi]))
    96       else
    97         (kind, phi)
    98   in
    99     "fof(" ^ ident ^ ", " ^ string_for_kind kind ^ ",\n    (" ^
   100     string_for_formula phi ^ ")" ^
   101     (case source of
   102        SOME tm => ", " ^ string_for_term tm
   103      | NONE => "") ^ ").\n"
   104   end
   105 fun tptp_strings_for_atp_problem use_conjecture_for_hypotheses problem =
   106   "% This file was generated by Isabelle (most likely Sledgehammer)\n\
   107   \% " ^ timestamp () ^ "\n" ::
   108   maps (fn (_, []) => []
   109          | (heading, lines) =>
   110            "\n% " ^ heading ^ " (" ^ string_of_int (length lines) ^ ")\n" ::
   111            map (string_for_problem_line use_conjecture_for_hypotheses) lines)
   112        problem
   113 
   114 fun is_atp_variable s = Char.isUpper (String.sub (s, 0))
   115 
   116 
   117 (** Nice names **)
   118 
   119 fun empty_name_pool readable_names =
   120   if readable_names then SOME (Symtab.empty, Symtab.empty) else NONE
   121 
   122 fun pool_fold f xs z = pair z #> fold_rev (fn x => uncurry (f x)) xs
   123 fun pool_map f xs =
   124   pool_fold (fn x => fn ys => fn pool => f x pool |>> (fn y => y :: ys)) xs []
   125 
   126 val no_qualifiers =
   127   let
   128     fun skip [] = []
   129       | skip (#"." :: cs) = skip cs
   130       | skip (c :: cs) = if Char.isAlphaNum c then skip cs else c :: keep cs
   131     and keep [] = []
   132       | keep (#"." :: cs) = skip cs
   133       | keep (c :: cs) = c :: keep cs
   134   in String.explode #> rev #> keep #> rev #> String.implode end
   135 
   136 (* "op" is also reserved, to avoid the unreadable "op_1", "op_2", etc., in the
   137    problem files. "equal" is reserved by some ATPs. "eq" is reserved to ensure
   138    that "HOL.eq" is correctly mapped to equality. *)
   139 val reserved_nice_names = ["op", "equal", "eq"]
   140 fun readable_name full_name s =
   141   if s = full_name then
   142     s
   143   else
   144     let
   145       val s = s |> no_qualifiers
   146                 |> Name.desymbolize (Char.isUpper (String.sub (full_name, 0)))
   147     in if member (op =) reserved_nice_names s then full_name else s end
   148 
   149 fun nice_name (full_name, _) NONE = (full_name, NONE)
   150   | nice_name (full_name, desired_name) (SOME the_pool) =
   151     if String.isPrefix "$" full_name then
   152       (full_name, SOME the_pool)
   153     else case Symtab.lookup (fst the_pool) full_name of
   154       SOME nice_name => (nice_name, SOME the_pool)
   155     | NONE =>
   156       let
   157         val nice_prefix = readable_name full_name desired_name
   158         fun add j =
   159           let
   160             val nice_name = nice_prefix ^
   161                             (if j = 0 then "" else "_" ^ string_of_int j)
   162           in
   163             case Symtab.lookup (snd the_pool) nice_name of
   164               SOME full_name' =>
   165               if full_name = full_name' then (nice_name, the_pool)
   166               else add (j + 1)
   167             | NONE =>
   168               (nice_name,
   169                (Symtab.update_new (full_name, nice_name) (fst the_pool),
   170                 Symtab.update_new (nice_name, full_name) (snd the_pool)))
   171           end
   172       in add 0 |> apsnd SOME end
   173 
   174 fun nice_term (ATerm (name, ts)) =
   175   nice_name name ##>> pool_map nice_term ts #>> ATerm
   176 fun nice_formula (AQuant (q, xs, phi)) =
   177     pool_map nice_name xs ##>> nice_formula phi
   178     #>> (fn (xs, phi) => AQuant (q, xs, phi))
   179   | nice_formula (AConn (c, phis)) =
   180     pool_map nice_formula phis #>> curry AConn c
   181   | nice_formula (AAtom tm) = nice_term tm #>> AAtom
   182 fun nice_problem_line (Fof (ident, kind, phi, source)) =
   183   nice_formula phi #>> (fn phi => Fof (ident, kind, phi, source))
   184 fun nice_problem problem =
   185   pool_map (fn (heading, lines) =>
   186                pool_map nice_problem_line lines #>> pair heading) problem
   187 fun nice_atp_problem readable_names problem =
   188   nice_problem problem (empty_name_pool readable_names)
   189 
   190 end;