src/HOL/Tools/ATP/atp_problem.ML
author blanchet
Sun, 13 May 2012 16:31:01 +0200
changeset 48926 2168126446bb
parent 48638 0b2b7ff31867
child 48991 6b13451135a9
permissions -rw-r--r--
extend ATP data structure to avoid having to perform ((non-)capture avoiding) beta reduction -- fixes a bug in the THF translation of "is_measure.simps"
     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, 'b) ho_term =
    11     ATerm of 'a * ('a, 'b) ho_term list |
    12     AAbs of (('a * 'b) * ('a, 'b) ho_term) * ('a, 'b) ho_term list
    13   datatype quantifier = AForall | AExists
    14   datatype connective = ANot | AAnd | AOr | AImplies | AIff
    15   datatype ('a, 'b, 'c) formula =
    16     AQuant of quantifier * ('a * 'b option) list * ('a, 'b, 'c) formula |
    17     AConn of connective * ('a, 'b, 'c) formula list |
    18     AAtom of 'c
    19 
    20   datatype 'a ho_type =
    21     AType of 'a * 'a ho_type list |
    22     AFun of 'a ho_type * 'a ho_type |
    23     ATyAbs of 'a list * 'a ho_type
    24 
    25   type term_order =
    26     {is_lpo : bool,
    27      gen_weights : bool,
    28      gen_prec : bool,
    29      gen_simp : bool}
    30 
    31   datatype tptp_polymorphism = TPTP_Monomorphic | TPTP_Polymorphic
    32   datatype tptp_explicitness = TPTP_Implicit | TPTP_Explicit
    33   datatype thf_flavor = THF_Without_Choice | THF_With_Choice
    34   datatype dfg_flavor = DFG_Unsorted | DFG_Sorted
    35 
    36   datatype atp_format =
    37     CNF |
    38     CNF_UEQ |
    39     FOF |
    40     TFF of tptp_polymorphism * tptp_explicitness |
    41     THF of tptp_polymorphism * tptp_explicitness * thf_flavor |
    42     DFG of dfg_flavor
    43 
    44   datatype formula_kind = Axiom | Definition | Lemma | Hypothesis | Conjecture
    45   datatype 'a problem_line =
    46     Decl of string * 'a * 'a ho_type |
    47     Formula of string * formula_kind
    48                * ('a, 'a ho_type, ('a, 'a ho_type) ho_term) formula
    49                * (string, string ho_type) ho_term option
    50                * (string, string ho_type) ho_term list
    51   type 'a problem = (string * 'a problem_line list) list
    52 
    53   val tptp_cnf : string
    54   val tptp_fof : string
    55   val tptp_tff : string
    56   val tptp_thf : string
    57   val tptp_has_type : string
    58   val tptp_type_of_types : string
    59   val tptp_bool_type : string
    60   val tptp_individual_type : string
    61   val tptp_fun_type : string
    62   val tptp_product_type : string
    63   val tptp_forall : string
    64   val tptp_ho_forall : string
    65   val tptp_pi_binder : string
    66   val tptp_exists : string
    67   val tptp_ho_exists : string
    68   val tptp_choice : string
    69   val tptp_not : string
    70   val tptp_and : string
    71   val tptp_or : string
    72   val tptp_implies : string
    73   val tptp_if : string
    74   val tptp_iff : string
    75   val tptp_not_iff : string
    76   val tptp_app : string
    77   val tptp_not_infix : string
    78   val tptp_equal : string
    79   val tptp_old_equal : string
    80   val tptp_false : string
    81   val tptp_true : string
    82   val tptp_empty_list : string
    83   val isabelle_info_prefix : string
    84   val isabelle_info : string -> int -> (string, 'a) ho_term list
    85   val extract_isabelle_status : (string, 'a) ho_term list -> string option
    86   val extract_isabelle_rank : (string, 'a) ho_term list -> int
    87   val introN : string
    88   val inductiveN : string
    89   val elimN : string
    90   val simpN : string
    91   val defN : string
    92   val rankN : string
    93   val minimum_rank : int
    94   val default_rank : int
    95   val default_term_order_weight : int
    96   val is_tptp_equal : string -> bool
    97   val is_built_in_tptp_symbol : string -> bool
    98   val is_tptp_variable : string -> bool
    99   val is_tptp_user_symbol : string -> bool
   100   val atype_of_types : (string * string) ho_type
   101   val bool_atype : (string * string) ho_type
   102   val individual_atype : (string * string) ho_type
   103   val mk_anot : ('a, 'b, 'c) formula -> ('a, 'b, 'c) formula
   104   val mk_aconn :
   105     connective -> ('a, 'b, 'c) formula -> ('a, 'b, 'c) formula
   106     -> ('a, 'b, 'c) formula
   107   val aconn_fold :
   108     bool option -> (bool option -> 'a -> 'b -> 'b) -> connective * 'a list
   109     -> 'b -> 'b
   110   val aconn_map :
   111     bool option -> (bool option -> 'a -> ('b, 'c, 'd) formula)
   112     -> connective * 'a list -> ('b, 'c, 'd) formula
   113   val formula_fold :
   114     bool option -> (bool option -> 'c -> 'd -> 'd) -> ('a, 'b, 'c) formula
   115     -> 'd -> 'd
   116   val formula_map : ('c -> 'd) -> ('a, 'b, 'c) formula -> ('a, 'b, 'd) formula
   117   val is_function_type : string ho_type -> bool
   118   val is_predicate_type : string ho_type -> bool
   119   val is_format_higher_order : atp_format -> bool
   120   val lines_for_atp_problem :
   121     atp_format -> term_order -> (unit -> (string * int) list) -> string problem
   122     -> string list
   123   val ensure_cnf_problem :
   124     (string * string) problem -> (string * string) problem
   125   val filter_cnf_ueq_problem :
   126     (string * string) problem -> (string * string) problem
   127   val declared_syms_in_problem : 'a problem -> 'a list
   128   val nice_atp_problem :
   129     bool -> atp_format -> ('a * (string * string) problem_line list) list
   130     -> ('a * string problem_line list) list
   131        * (string Symtab.table * string Symtab.table) option
   132 end;
   133 
   134 structure ATP_Problem : ATP_PROBLEM =
   135 struct
   136 
   137 open ATP_Util
   138 
   139 
   140 (** ATP problem **)
   141 
   142 datatype ('a, 'b) ho_term =
   143   ATerm of 'a * ('a, 'b) ho_term list |
   144   AAbs of (('a * 'b) * ('a, 'b) ho_term) * ('a, 'b) ho_term list
   145 datatype quantifier = AForall | AExists
   146 datatype connective = ANot | AAnd | AOr | AImplies | AIff
   147 datatype ('a, 'b, 'c) formula =
   148   AQuant of quantifier * ('a * 'b option) list * ('a, 'b, 'c) formula |
   149   AConn of connective * ('a, 'b, 'c) formula list |
   150   AAtom of 'c
   151 
   152 datatype 'a ho_type =
   153   AType of 'a * 'a ho_type list |
   154   AFun of 'a ho_type * 'a ho_type |
   155   ATyAbs of 'a list * 'a ho_type
   156 
   157 type term_order =
   158   {is_lpo : bool,
   159    gen_weights : bool,
   160    gen_prec : bool,
   161    gen_simp : bool}
   162 
   163 datatype tptp_polymorphism = TPTP_Monomorphic | TPTP_Polymorphic
   164 datatype tptp_explicitness = TPTP_Implicit | TPTP_Explicit
   165 datatype thf_flavor = THF_Without_Choice | THF_With_Choice
   166 datatype dfg_flavor = DFG_Unsorted | DFG_Sorted
   167 
   168 datatype atp_format =
   169   CNF |
   170   CNF_UEQ |
   171   FOF |
   172   TFF of tptp_polymorphism * tptp_explicitness |
   173   THF of tptp_polymorphism * tptp_explicitness * thf_flavor |
   174   DFG of dfg_flavor
   175 
   176 datatype formula_kind = Axiom | Definition | Lemma | Hypothesis | Conjecture
   177 datatype 'a problem_line =
   178   Decl of string * 'a * 'a ho_type |
   179   Formula of string * formula_kind
   180              * ('a, 'a ho_type, ('a, 'a ho_type) ho_term) formula
   181              * (string, string ho_type) ho_term option
   182              * (string, string ho_type) ho_term list
   183 type 'a problem = (string * 'a problem_line list) list
   184 
   185 (* official TPTP syntax *)
   186 val tptp_cnf = "cnf"
   187 val tptp_fof = "fof"
   188 val tptp_tff = "tff"
   189 val tptp_thf = "thf"
   190 val tptp_has_type = ":"
   191 val tptp_type_of_types = "$tType"
   192 val tptp_bool_type = "$o"
   193 val tptp_individual_type = "$i"
   194 val tptp_fun_type = ">"
   195 val tptp_product_type = "*"
   196 val tptp_forall = "!"
   197 val tptp_ho_forall = "!!"
   198 val tptp_pi_binder = "!>"
   199 val tptp_exists = "?"
   200 val tptp_ho_exists = "??"
   201 val tptp_choice = "@+"
   202 val tptp_not = "~"
   203 val tptp_and = "&"
   204 val tptp_or = "|"
   205 val tptp_implies = "=>"
   206 val tptp_if = "<="
   207 val tptp_iff = "<=>"
   208 val tptp_not_iff = "<~>"
   209 val tptp_app = "@"
   210 val tptp_not_infix = "!"
   211 val tptp_equal = "="
   212 val tptp_old_equal = "equal"
   213 val tptp_false = "$false"
   214 val tptp_true = "$true"
   215 val tptp_empty_list = "[]"
   216 
   217 val isabelle_info_prefix = "isabelle_"
   218 
   219 val introN = "intro"
   220 val inductiveN = "inductive"
   221 val elimN = "elim"
   222 val simpN = "simp"
   223 val defN = "def"
   224 val rankN = "rank"
   225 
   226 val minimum_rank = 0
   227 val default_rank = 1000
   228 val default_term_order_weight = 1
   229 
   230 (* Currently, only newer versions of SPASS, with sorted DFG format support, can
   231    process Isabelle metainformation. *)
   232 fun isabelle_info status rank =
   233   [] |> rank <> default_rank
   234         ? cons (ATerm (isabelle_info_prefix ^ rankN,
   235                        [ATerm (string_of_int rank, [])]))
   236      |> status <> "" ? cons (ATerm (isabelle_info_prefix ^ status, []))
   237 
   238 fun extract_isabelle_status (ATerm (s, []) :: _) =
   239     try (unprefix isabelle_info_prefix) s
   240   | extract_isabelle_status _ = NONE
   241 
   242 fun extract_isabelle_rank (tms as _ :: _) =
   243     (case List.last tms of
   244        ATerm (_, [ATerm (rank, [])]) => the (Int.fromString rank)
   245      | _ => default_rank)
   246   | extract_isabelle_rank _ = default_rank
   247 
   248 fun is_tptp_equal s = (s = tptp_equal orelse s = tptp_old_equal)
   249 fun is_built_in_tptp_symbol s =
   250   s = tptp_old_equal orelse not (Char.isAlpha (String.sub (s, 0)))
   251 fun is_tptp_variable s = Char.isUpper (String.sub (s, 0))
   252 val is_tptp_user_symbol = not o (is_tptp_variable orf is_built_in_tptp_symbol)
   253 
   254 val atype_of_types = AType (`I tptp_type_of_types, [])
   255 val bool_atype = AType (`I tptp_bool_type, [])
   256 val individual_atype = AType (`I tptp_individual_type, [])
   257 
   258 fun raw_polarities_of_conn ANot = (SOME false, NONE)
   259   | raw_polarities_of_conn AAnd = (SOME true, SOME true)
   260   | raw_polarities_of_conn AOr = (SOME true, SOME true)
   261   | raw_polarities_of_conn AImplies = (SOME false, SOME true)
   262   | raw_polarities_of_conn AIff = (NONE, NONE)
   263 fun polarities_of_conn NONE = K (NONE, NONE)
   264   | polarities_of_conn (SOME pos) =
   265     raw_polarities_of_conn #> not pos ? pairself (Option.map not)
   266 
   267 fun mk_anot (AConn (ANot, [phi])) = phi
   268   | mk_anot phi = AConn (ANot, [phi])
   269 fun mk_aconn c phi1 phi2 = AConn (c, [phi1, phi2])
   270 
   271 fun aconn_fold pos f (ANot, [phi]) = f (Option.map not pos) phi
   272   | aconn_fold pos f (AImplies, [phi1, phi2]) =
   273     f (Option.map not pos) phi1 #> f pos phi2
   274   | aconn_fold pos f (AAnd, phis) = fold (f pos) phis
   275   | aconn_fold pos f (AOr, phis) = fold (f pos) phis
   276   | aconn_fold _ f (_, phis) = fold (f NONE) phis
   277 
   278 fun aconn_map pos f (ANot, [phi]) = AConn (ANot, [f (Option.map not pos) phi])
   279   | aconn_map pos f (AImplies, [phi1, phi2]) =
   280     AConn (AImplies, [f (Option.map not pos) phi1, f pos phi2])
   281   | aconn_map pos f (AAnd, phis) = AConn (AAnd, map (f pos) phis)
   282   | aconn_map pos f (AOr, phis) = AConn (AOr, map (f pos) phis)
   283   | aconn_map _ f (c, phis) = AConn (c, map (f NONE) phis)
   284 
   285 fun formula_fold pos f =
   286   let
   287     fun fld pos (AQuant (_, _, phi)) = fld pos phi
   288       | fld pos (AConn conn) = aconn_fold pos fld conn
   289       | fld pos (AAtom tm) = f pos tm
   290   in fld pos end
   291 
   292 fun formula_map f (AQuant (q, xs, phi)) = AQuant (q, xs, formula_map f phi)
   293   | formula_map f (AConn (c, phis)) = AConn (c, map (formula_map f) phis)
   294   | formula_map f (AAtom tm) = AAtom (f tm)
   295 
   296 fun is_function_type (AFun (_, ty)) = is_function_type ty
   297   | is_function_type (AType (s, _)) =
   298     s <> tptp_type_of_types andalso s <> tptp_bool_type
   299   | is_function_type _ = false
   300 fun is_predicate_type (AFun (_, ty)) = is_predicate_type ty
   301   | is_predicate_type (AType (s, _)) = (s = tptp_bool_type)
   302   | is_predicate_type _ = false
   303 fun is_nontrivial_predicate_type (AFun (_, ty)) = is_predicate_type ty
   304   | is_nontrivial_predicate_type _ = false
   305 
   306 fun is_format_higher_order (THF _) = true
   307   | is_format_higher_order _ = false
   308 fun is_format_typed (TFF _) = true
   309   | is_format_typed (THF _) = true
   310   | is_format_typed (DFG DFG_Sorted) = true
   311   | is_format_typed _ = false
   312 
   313 fun tptp_string_for_kind Axiom = "axiom"
   314   | tptp_string_for_kind Definition = "definition"
   315   | tptp_string_for_kind Lemma = "lemma"
   316   | tptp_string_for_kind Hypothesis = "hypothesis"
   317   | tptp_string_for_kind Conjecture = "conjecture"
   318 
   319 fun tptp_string_for_app format func args =
   320   if is_format_higher_order format then
   321     "(" ^ space_implode (" " ^ tptp_app ^ " ") (func :: args) ^ ")"
   322   else
   323     func ^ "(" ^ commas args ^ ")"
   324 
   325 fun flatten_type (ATyAbs (tys, ty)) = ATyAbs (tys, flatten_type ty)
   326   | flatten_type (ty as AFun (ty1 as AType _, ty2)) =
   327     (case flatten_type ty2 of
   328        AFun (ty' as AType (s, tys), ty) =>
   329        AFun (AType (tptp_product_type,
   330                     ty1 :: (if s = tptp_product_type then tys else [ty'])), ty)
   331      | _ => ty)
   332   | flatten_type (ty as AType _) = ty
   333   | flatten_type _ =
   334     raise Fail "unexpected higher-order type in first-order format"
   335 
   336 val dfg_individual_type = "iii" (* cannot clash *)
   337 
   338 fun str_for_type format ty =
   339   let
   340     val dfg = (format = DFG DFG_Sorted)
   341     fun str _ (AType (s, [])) =
   342         if dfg andalso s = tptp_individual_type then dfg_individual_type else s
   343       | str _ (AType (s, tys)) =
   344         let val ss = tys |> map (str false) in
   345           if s = tptp_product_type then
   346             ss |> space_implode
   347                       (if dfg then ", " else " " ^ tptp_product_type ^ " ")
   348                |> (not dfg andalso length ss > 1) ? enclose "(" ")"
   349           else
   350             tptp_string_for_app format s ss
   351         end
   352       | str rhs (AFun (ty1, ty2)) =
   353         (str false ty1 |> dfg ? enclose "(" ")") ^ " " ^
   354         (if dfg then "" else tptp_fun_type ^ " ") ^ str true ty2
   355         |> not rhs ? enclose "(" ")"
   356       | str _ (ATyAbs (ss, ty)) =
   357         tptp_pi_binder ^ "[" ^
   358         commas (map (suffix (" " ^ tptp_has_type ^ " " ^ tptp_type_of_types))
   359                     ss) ^ "]: " ^ str false ty
   360   in str true ty end
   361 
   362 fun string_for_type (format as THF _) ty = str_for_type format ty
   363   | string_for_type format ty = str_for_type format (flatten_type ty)
   364 
   365 fun tptp_string_for_quantifier AForall = tptp_forall
   366   | tptp_string_for_quantifier AExists = tptp_exists
   367 
   368 fun tptp_string_for_connective ANot = tptp_not
   369   | tptp_string_for_connective AAnd = tptp_and
   370   | tptp_string_for_connective AOr = tptp_or
   371   | tptp_string_for_connective AImplies = tptp_implies
   372   | tptp_string_for_connective AIff = tptp_iff
   373 
   374 fun string_for_bound_var format (s, ty) =
   375   s ^
   376   (if is_format_typed format then
   377      " " ^ tptp_has_type ^ " " ^
   378      (ty |> the_default (AType (tptp_individual_type, []))
   379          |> string_for_type format)
   380    else
   381      "")
   382 
   383 fun is_format_with_choice (THF (_, _, THF_With_Choice)) = true
   384   | is_format_with_choice _ = false
   385 
   386 fun tptp_string_for_term _ (ATerm (s, [])) = s
   387   | tptp_string_for_term format (ATerm (s, ts)) =
   388     (if s = tptp_empty_list then
   389        (* used for lists in the optional "source" field of a derivation *)
   390        "[" ^ commas (map (tptp_string_for_term format) ts) ^ "]"
   391      else if is_tptp_equal s then
   392        space_implode (" " ^ tptp_equal ^ " ")
   393                      (map (tptp_string_for_term format) ts)
   394        |> is_format_higher_order format ? enclose "(" ")"
   395      else case (s = tptp_ho_forall orelse s = tptp_ho_exists,
   396                 s = tptp_choice andalso is_format_with_choice format, ts) of
   397        (true, _, [AAbs (((s', ty), tm), [])]) =>
   398        (* Use syntactic sugar "!" and "?" instead of "!!" and "??" whenever
   399           possible, to work around LEO-II 1.2.8 parser limitation. *)
   400        tptp_string_for_formula format
   401            (AQuant (if s = tptp_ho_forall then AForall else AExists,
   402                     [(s', SOME ty)], AAtom tm))
   403      | (_, true, [AAbs (((s', ty), tm), args)]) =>
   404        (* There is code in "ATP_Problem_Generate" to ensure that "Eps" is always
   405           applied to an abstraction. *)
   406        tptp_string_for_app format
   407            (tptp_choice ^ "[" ^ s' ^ " : " ^ string_for_type format ty ^ "]: " ^
   408             tptp_string_for_term format tm ^ ""
   409             |> enclose "(" ")")
   410            (map (tptp_string_for_term format) args)
   411      | _ => tptp_string_for_app format s (map (tptp_string_for_term format) ts))
   412   | tptp_string_for_term (format as THF _) (AAbs (((s, ty), tm), args)) =
   413     tptp_string_for_app format
   414         ("(^[" ^ s ^ " : " ^ string_for_type format ty ^ "]: " ^
   415          tptp_string_for_term format tm ^ ")")
   416         (map (tptp_string_for_term format) args)
   417   | tptp_string_for_term _ _ =
   418     raise Fail "unexpected term in first-order format"
   419 and tptp_string_for_formula format (AQuant (q, xs, phi)) =
   420     tptp_string_for_quantifier q ^
   421     "[" ^ commas (map (string_for_bound_var format) xs) ^ "]: " ^
   422     tptp_string_for_formula format phi
   423     |> enclose "(" ")"
   424   | tptp_string_for_formula format
   425         (AConn (ANot, [AAtom (ATerm ("=" (* tptp_equal *), ts))])) =
   426     space_implode (" " ^ tptp_not_infix ^ tptp_equal ^ " ")
   427                   (map (tptp_string_for_term format) ts)
   428     |> is_format_higher_order format ? enclose "(" ")"
   429   | tptp_string_for_formula format (AConn (c, [phi])) =
   430     tptp_string_for_connective c ^ " " ^
   431     (tptp_string_for_formula format phi
   432      |> is_format_higher_order format ? enclose "(" ")")
   433     |> enclose "(" ")"
   434   | tptp_string_for_formula format (AConn (c, phis)) =
   435     space_implode (" " ^ tptp_string_for_connective c ^ " ")
   436                   (map (tptp_string_for_formula format) phis)
   437     |> enclose "(" ")"
   438   | tptp_string_for_formula format (AAtom tm) = tptp_string_for_term format tm
   439 
   440 fun tptp_string_for_format CNF = tptp_cnf
   441   | tptp_string_for_format CNF_UEQ = tptp_cnf
   442   | tptp_string_for_format FOF = tptp_fof
   443   | tptp_string_for_format (TFF _) = tptp_tff
   444   | tptp_string_for_format (THF _) = tptp_thf
   445   | tptp_string_for_format (DFG _) = raise Fail "non-TPTP format"
   446 
   447 fun tptp_string_for_problem_line format (Decl (ident, sym, ty)) =
   448     tptp_string_for_format format ^ "(" ^ ident ^ ", type,\n    " ^ sym ^
   449     " : " ^ string_for_type format ty ^ ").\n"
   450   | tptp_string_for_problem_line format
   451                                  (Formula (ident, kind, phi, source, _)) =
   452     tptp_string_for_format format ^ "(" ^ ident ^ ", " ^
   453     tptp_string_for_kind kind ^ ",\n    (" ^
   454     tptp_string_for_formula format phi ^ ")" ^
   455     (case source of
   456        SOME tm => ", " ^ tptp_string_for_term format tm
   457      | NONE => "") ^ ").\n"
   458 
   459 fun tptp_lines format =
   460   maps (fn (_, []) => []
   461          | (heading, lines) =>
   462            "\n% " ^ heading ^ " (" ^ string_of_int (length lines) ^ ")\n" ::
   463            map (tptp_string_for_problem_line format) lines)
   464 
   465 fun arity_of_type (AFun (_, ty)) = 1 + arity_of_type ty
   466   | arity_of_type _ = 0
   467 
   468 fun binder_atypes (AFun (ty1, ty2)) = ty1 :: binder_atypes ty2
   469   | binder_atypes _ = []
   470 
   471 fun dfg_string_for_formula gen_simp flavor info =
   472   let
   473     fun suffix_tag top_level s =
   474       if flavor = DFG_Sorted andalso top_level then
   475         case extract_isabelle_status info of
   476           SOME s' => if s' = defN then s ^ ":lt"
   477                      else if s' = simpN andalso gen_simp then s ^ ":lr"
   478                      else s
   479         | NONE => s
   480       else
   481         s
   482     fun str_for_term top_level (ATerm (s, tms)) =
   483         (if is_tptp_equal s then "equal" |> suffix_tag top_level
   484          else if s = tptp_true then "true"
   485          else if s = tptp_false then "false"
   486          else s) ^
   487         (if null tms then ""
   488          else "(" ^ commas (map (str_for_term false) tms) ^ ")")
   489       | str_for_term _ _ = raise Fail "unexpected term in first-order format"
   490     fun str_for_quant AForall = "forall"
   491       | str_for_quant AExists = "exists"
   492     fun str_for_conn _ ANot = "not"
   493       | str_for_conn _ AAnd = "and"
   494       | str_for_conn _ AOr = "or"
   495       | str_for_conn _ AImplies = "implies"
   496       | str_for_conn top_level AIff = "equiv" |> suffix_tag top_level
   497     fun str_for_formula top_level (AQuant (q, xs, phi)) =
   498         str_for_quant q ^ "(" ^ "[" ^
   499         commas (map (string_for_bound_var (DFG flavor)) xs) ^ "], " ^
   500         str_for_formula top_level phi ^ ")"
   501       | str_for_formula top_level (AConn (c, phis)) =
   502         str_for_conn top_level c ^ "(" ^
   503         commas (map (str_for_formula false) phis) ^ ")"
   504       | str_for_formula top_level (AAtom tm) = str_for_term top_level tm
   505   in str_for_formula true end
   506 
   507 fun maybe_enclose bef aft "" = "% " ^ bef ^ aft
   508   | maybe_enclose bef aft s = bef ^ s ^ aft
   509 
   510 fun dfg_lines flavor {is_lpo, gen_weights, gen_prec, gen_simp} ord_info
   511               problem =
   512   let
   513     val sorted = (flavor = DFG_Sorted)
   514     val format = DFG flavor
   515     fun spair (sym, k) = "(" ^ sym ^ ", " ^ string_of_int k ^ ")"
   516     fun ary sym = curry spair sym o arity_of_type
   517     fun fun_typ sym ty =
   518       "function(" ^ sym ^ ", " ^ string_for_type format ty ^ ")."
   519     fun pred_typ sym ty =
   520       "predicate(" ^
   521       commas (sym :: map (string_for_type format) (binder_atypes ty)) ^ ")."
   522     fun formula pred (Formula (ident, kind, phi, _, info)) =
   523         if pred kind then
   524           let
   525             val rank =
   526               if flavor = DFG_Sorted then extract_isabelle_rank info
   527               else default_rank
   528           in
   529             "formula(" ^ dfg_string_for_formula gen_simp flavor info phi ^
   530             ", " ^ ident ^
   531             (if rank = default_rank then "" else ", " ^ string_of_int rank) ^
   532             ")." |> SOME
   533           end
   534         else
   535           NONE
   536       | formula _ _ = NONE
   537     fun filt f = problem |> map (map_filter f o snd) |> filter_out null
   538     val func_aries =
   539       filt (fn Decl (_, sym, ty) =>
   540                if is_function_type ty then SOME (ary sym ty) else NONE
   541              | _ => NONE)
   542       |> flat |> commas |> maybe_enclose "functions [" "]."
   543     val pred_aries =
   544       filt (fn Decl (_, sym, ty) =>
   545                if is_predicate_type ty then SOME (ary sym ty) else NONE
   546              | _ => NONE)
   547       |> flat |> commas |> maybe_enclose "predicates [" "]."
   548     fun sorts () =
   549       filt (fn Decl (_, sym, AType (s, [])) =>
   550                if s = tptp_type_of_types then SOME sym else NONE
   551              | _ => NONE) @ [[dfg_individual_type]]
   552       |> flat |> commas |> maybe_enclose "sorts [" "]."
   553     val ord_info =
   554       if sorted andalso (gen_weights orelse gen_prec) then ord_info () else []
   555     fun do_term_order_weights () =
   556       (if gen_weights then ord_info else [])
   557       |> map spair |> commas |> maybe_enclose "weights [" "]."
   558     val syms =
   559       [func_aries, pred_aries] @
   560       (if sorted then [do_term_order_weights (), sorts ()] else [])
   561     fun func_sigs () =
   562       filt (fn Decl (_, sym, ty) =>
   563                if is_function_type ty then SOME (fun_typ sym ty) else NONE
   564              | _ => NONE)
   565       |> flat
   566     fun pred_sigs () =
   567       filt (fn Decl (_, sym, ty) =>
   568                if is_nontrivial_predicate_type ty then SOME (pred_typ sym ty)
   569                else NONE
   570              | _ => NONE)
   571       |> flat
   572     val decls = if sorted then func_sigs () @ pred_sigs () else []
   573     val axioms =
   574       filt (formula (curry (op <>) Conjecture)) |> separate [""] |> flat
   575     val conjs =
   576       filt (formula (curry (op =) Conjecture)) |> separate [""] |> flat
   577     val settings =
   578       (if is_lpo then ["set_flag(Ordering, 1)."] else []) @
   579       (if gen_prec then
   580          [ord_info |> map fst |> rev |> commas
   581                    |> maybe_enclose "set_precedence(" ")."]
   582        else
   583          [])
   584     fun list_of _ [] = []
   585       | list_of heading ss =
   586         "list_of_" ^ heading ^ ".\n" :: map (suffix "\n") ss @
   587         ["end_of_list.\n\n"]
   588   in
   589     "\nbegin_problem(isabelle).\n\n" ::
   590     list_of "descriptions"
   591             ["name({**}).", "author({**}).", "status(unknown).",
   592              "description({**})."] @
   593     list_of "symbols" syms @
   594     list_of "declarations" decls @
   595     list_of "formulae(axioms)" axioms @
   596     list_of "formulae(conjectures)" conjs @
   597     list_of "settings(SPASS)" settings @
   598     ["end_problem.\n"]
   599   end
   600 
   601 fun lines_for_atp_problem format ord ord_info problem =
   602   "% This file was generated by Isabelle (most likely Sledgehammer)\n\
   603   \% " ^ timestamp () ^ "\n" ::
   604   (case format of
   605      DFG flavor => dfg_lines flavor ord ord_info
   606    | _ => tptp_lines format) problem
   607 
   608 
   609 (** CNF (Metis) and CNF UEQ (Waldmeister) **)
   610 
   611 fun is_problem_line_negated (Formula (_, _, AConn (ANot, _), _, _)) = true
   612   | is_problem_line_negated _ = false
   613 
   614 fun is_problem_line_cnf_ueq (Formula (_, _, AAtom (ATerm ((s, _), _)), _, _)) =
   615     is_tptp_equal s
   616   | is_problem_line_cnf_ueq _ = false
   617 
   618 fun open_conjecture_term (ATerm ((s, s'), tms)) =
   619     ATerm (if is_tptp_variable s then (s |> Name.desymbolize false, s')
   620            else (s, s'), tms |> map open_conjecture_term)
   621   | open_conjecture_term _ = raise Fail "unexpected higher-order term"
   622 fun open_formula conj =
   623   let
   624     (* We are conveniently assuming that all bound variable names are
   625        distinct, which should be the case for the formulas we generate. *)
   626     fun opn (pos as SOME true) (AQuant (AForall, _, phi)) = opn pos phi
   627       | opn (pos as SOME false) (AQuant (AExists, _, phi)) = opn pos phi
   628       | opn pos (AConn (ANot, [phi])) = mk_anot (opn (Option.map not pos) phi)
   629       | opn pos (AConn (c, [phi1, phi2])) =
   630         let val (pos1, pos2) = polarities_of_conn pos c in
   631           AConn (c, [opn pos1 phi1, opn pos2 phi2])
   632         end
   633       | opn _ (AAtom t) = AAtom (t |> conj ? open_conjecture_term)
   634       | opn _ phi = phi
   635   in opn (SOME (not conj)) end
   636 fun open_formula_line (Formula (ident, kind, phi, source, info)) =
   637     Formula (ident, kind, open_formula (kind = Conjecture) phi, source, info)
   638   | open_formula_line line = line
   639 
   640 fun negate_conjecture_line (Formula (ident, Conjecture, phi, source, info)) =
   641     Formula (ident, Hypothesis, mk_anot phi, source, info)
   642   | negate_conjecture_line line = line
   643 
   644 exception CLAUSIFY of unit
   645 
   646 (* This "clausification" only expands syntactic sugar, such as "phi => psi" to
   647    "~ phi | psi" and "phi <=> psi" to "~ phi | psi" and "~ psi | phi". We don't
   648    attempt to distribute conjunctions over disjunctions. *)
   649 fun clausify_formula pos (phi as AAtom _) = [phi |> not pos ? mk_anot]
   650   | clausify_formula pos (AConn (ANot, [phi])) = clausify_formula (not pos) phi
   651   | clausify_formula true (AConn (AOr, [phi1, phi2])) =
   652     (phi1, phi2) |> pairself (clausify_formula true)
   653                  |> uncurry (map_product (mk_aconn AOr))
   654   | clausify_formula false (AConn (AAnd, [phi1, phi2])) =
   655     (phi1, phi2) |> pairself (clausify_formula false)
   656                  |> uncurry (map_product (mk_aconn AOr))
   657   | clausify_formula true (AConn (AImplies, [phi1, phi2])) =
   658     clausify_formula true (AConn (AOr, [mk_anot phi1, phi2]))
   659   | clausify_formula true (AConn (AIff, phis)) =
   660     clausify_formula true (AConn (AImplies, phis)) @
   661     clausify_formula true (AConn (AImplies, rev phis))
   662   | clausify_formula _ _ = raise CLAUSIFY ()
   663 
   664 fun clausify_formula_line (Formula (ident, kind, phi, source, info)) =
   665     let
   666       val (n, phis) = phi |> try (clausify_formula true) |> these |> `length
   667     in
   668       map2 (fn phi => fn j =>
   669                Formula (ident ^ replicate_string (j - 1) "x", kind, phi, source,
   670                         info))
   671            phis (1 upto n)
   672     end
   673   | clausify_formula_line _ = []
   674 
   675 fun ensure_cnf_problem_line line =
   676   line |> open_formula_line |> negate_conjecture_line |> clausify_formula_line
   677 
   678 fun ensure_cnf_problem problem =
   679   problem |> map (apsnd (maps ensure_cnf_problem_line))
   680 
   681 fun filter_cnf_ueq_problem problem =
   682   problem
   683   |> map (apsnd (map open_formula_line
   684                  #> filter is_problem_line_cnf_ueq
   685                  #> map negate_conjecture_line))
   686   |> (fn problem =>
   687          let
   688            val lines = problem |> maps snd
   689            val conjs = lines |> filter is_problem_line_negated
   690          in if length conjs = 1 andalso conjs <> lines then problem else [] end)
   691 
   692 
   693 (** Symbol declarations **)
   694 
   695 fun add_declared_syms_in_problem_line (Decl (_, sym, _)) = cons sym
   696   | add_declared_syms_in_problem_line _ = I
   697 fun declared_syms_in_problem problem =
   698   fold (fold add_declared_syms_in_problem_line o snd) problem []
   699 
   700 (** Nice names **)
   701 
   702 fun pool_fold f xs z = pair z #> fold_rev (fn x => uncurry (f x)) xs
   703 fun pool_map f xs =
   704   pool_fold (fn x => fn ys => fn pool => f x pool |>> (fn y => y :: ys)) xs []
   705 
   706 val no_qualifiers =
   707   let
   708     fun skip [] = []
   709       | skip (#"." :: cs) = skip cs
   710       | skip (c :: cs) = if Char.isAlphaNum c then skip cs else c :: keep cs
   711     and keep [] = []
   712       | keep (#"." :: cs) = skip cs
   713       | keep (c :: cs) = c :: keep cs
   714   in String.explode #> rev #> keep #> rev #> String.implode end
   715 
   716 (* Long names can slow down the ATPs. *)
   717 val max_readable_name_size = 20
   718 
   719 (* "equal" is reserved by some ATPs. "op" is also reserved, to avoid the
   720    unreadable "op_1", "op_2", etc., in the problem files. "eq" is reserved to
   721    ensure that "HOL.eq" is correctly mapped to equality (not clear whether this
   722    is still necessary). *)
   723 val reserved_nice_names = [tptp_old_equal, "op", "eq"]
   724 
   725 (* hack to get the same hashing across Mirabelle runs (see "mirabelle.pl") *)
   726 fun cleanup_mirabelle_name s =
   727   let
   728     val mirabelle_infix = "_Mirabelle_"
   729     val random_suffix_len = 10
   730     val (s1, s2) = Substring.position mirabelle_infix (Substring.full s)
   731   in
   732     if Substring.isEmpty s2 then
   733       s
   734     else
   735       Substring.string s1 ^
   736       Substring.string (Substring.triml (size mirabelle_infix + random_suffix_len) s2)
   737   end
   738 
   739 fun readable_name protect full_name s =
   740   (if s = full_name then
   741      s
   742    else
   743      s |> no_qualifiers
   744        |> perhaps (try (unprefix "'"))
   745        |> Name.desymbolize (Char.isUpper (String.sub (full_name, 0)))
   746        |> (fn s =>
   747               if size s > max_readable_name_size then
   748                 String.substring (s, 0, max_readable_name_size div 2 - 4) ^
   749                 string_of_int (hash_string (cleanup_mirabelle_name full_name)) ^
   750                 String.extract (s, size s - max_readable_name_size div 2 + 4,
   751                                 NONE)
   752               else
   753                 s)
   754        |> (fn s =>
   755               if member (op =) reserved_nice_names s then full_name else s))
   756   |> protect
   757 
   758 fun nice_name _ (full_name, _) NONE = (full_name, NONE)
   759   | nice_name protect (full_name, desired_name) (SOME the_pool) =
   760     if is_built_in_tptp_symbol full_name then
   761       (full_name, SOME the_pool)
   762     else case Symtab.lookup (fst the_pool) full_name of
   763       SOME nice_name => (nice_name, SOME the_pool)
   764     | NONE =>
   765       let
   766         val nice_prefix = readable_name protect full_name desired_name
   767         fun add j =
   768           let
   769             val nice_name =
   770               nice_prefix ^ (if j = 1 then "" else string_of_int j)
   771           in
   772             case Symtab.lookup (snd the_pool) nice_name of
   773               SOME full_name' =>
   774               if full_name = full_name' then (nice_name, the_pool)
   775               else add (j + 1)
   776             | NONE =>
   777               (nice_name,
   778                (Symtab.update_new (full_name, nice_name) (fst the_pool),
   779                 Symtab.update_new (nice_name, full_name) (snd the_pool)))
   780           end
   781       in add 1 |> apsnd SOME end
   782 
   783 fun avoid_clash_with_alt_ergo_type_vars s =
   784   if is_tptp_variable s then s else s ^ "_"
   785 
   786 fun avoid_clash_with_dfg_keywords s =
   787   let val n = String.size s in
   788     if n < 2 orelse (n = 2 andalso String.sub (s, 0) = String.sub (s, 1)) orelse
   789        String.isSubstring "_" s then
   790       s
   791     else if is_tptp_variable s then
   792       (* "DL" appears to be a SPASS 3.7 keyword *)
   793       if s = "DL" then s ^ "_" else s
   794     else
   795       String.substring (s, 0, n - 1) ^
   796       String.str (Char.toUpper (String.sub (s, n - 1)))
   797   end
   798 
   799 fun nice_atp_problem readable_names format problem =
   800   let
   801     val empty_pool =
   802       if readable_names then SOME (Symtab.empty, Symtab.empty) else NONE
   803     val avoid_clash =
   804       case format of
   805         TFF (TPTP_Polymorphic, _) => avoid_clash_with_alt_ergo_type_vars
   806       | DFG _ => avoid_clash_with_dfg_keywords
   807       | _ => I
   808     val nice_name = nice_name avoid_clash
   809     fun nice_type (AType (name, tys)) =
   810         nice_name name ##>> pool_map nice_type tys #>> AType
   811       | nice_type (AFun (ty1, ty2)) = nice_type ty1 ##>> nice_type ty2 #>> AFun
   812       | nice_type (ATyAbs (names, ty)) =
   813         pool_map nice_name names ##>> nice_type ty #>> ATyAbs
   814     fun nice_term (ATerm (name, ts)) =
   815         nice_name name ##>> pool_map nice_term ts #>> ATerm
   816       | nice_term (AAbs (((name, ty), tm), args)) =
   817         nice_name name ##>> nice_type ty ##>> nice_term tm
   818         ##>> pool_map nice_term args #>> AAbs
   819     fun nice_formula (AQuant (q, xs, phi)) =
   820         pool_map nice_name (map fst xs)
   821         ##>> pool_map (fn NONE => pair NONE
   822                         | SOME ty => nice_type ty #>> SOME) (map snd xs)
   823         ##>> nice_formula phi
   824         #>> (fn ((ss, ts), phi) => AQuant (q, ss ~~ ts, phi))
   825       | nice_formula (AConn (c, phis)) =
   826         pool_map nice_formula phis #>> curry AConn c
   827       | nice_formula (AAtom tm) = nice_term tm #>> AAtom
   828     fun nice_problem_line (Decl (ident, sym, ty)) =
   829         nice_name sym ##>> nice_type ty #>> (fn (sym, ty) => Decl (ident, sym, ty))
   830       | nice_problem_line (Formula (ident, kind, phi, source, info)) =
   831         nice_formula phi #>> (fn phi => Formula (ident, kind, phi, source, info))
   832     fun nice_problem problem =
   833       pool_map (fn (heading, lines) =>
   834                    pool_map nice_problem_line lines #>> pair heading) problem
   835   in nice_problem problem empty_pool end
   836 
   837 end;