src/HOL/TPTP/atp_theory_export.ML
author blanchet
Tue, 10 Jul 2012 23:36:03 +0200
changeset 49248 50e00ee405f8
parent 49244 141ab3c13ac8
child 49249 06216c789ac9
permissions -rw-r--r--
distinguish updates and queries + cleanups
blanchet@47149
     1
(*  Title:      HOL/TPTP/atp_theory_export.ML
blanchet@43473
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@43473
     3
    Copyright   2011
blanchet@43473
     4
blanchet@49228
     5
Export Isabelle theories as MaSh (Machine-learning for Sledgehammer) or as
blanchet@49228
     6
first-order TPTP inferences.
blanchet@43473
     7
*)
blanchet@43473
     8
blanchet@47149
     9
signature ATP_THEORY_EXPORT =
blanchet@43473
    10
sig
blanchet@46176
    11
  type atp_format = ATP_Problem.atp_format
blanchet@46176
    12
blanchet@44365
    13
  val theorems_mentioned_in_proof_term :
blanchet@44365
    14
    string list option -> thm -> string list
blanchet@49231
    15
  val generate_mash_accessibility_file_for_theory :
blanchet@49231
    16
    theory -> bool -> string -> unit
blanchet@49231
    17
  val generate_mash_feature_file_for_theory : theory -> bool -> string -> unit
blanchet@49231
    18
  val generate_mash_dependency_file_for_theory :
blanchet@49231
    19
    theory -> bool -> string -> unit
blanchet@49231
    20
  val generate_mash_problem_file_for_theory : theory -> string -> unit
blanchet@43473
    21
  val generate_tptp_inference_file_for_theory :
blanchet@46176
    22
    Proof.context -> theory -> atp_format -> string -> string -> unit
blanchet@43473
    23
end;
blanchet@43473
    24
blanchet@49232
    25
structure ATP_Theory_Export (* ### : ATP_THEORY_EXPORT *) =
blanchet@43473
    26
struct
blanchet@43473
    27
blanchet@44350
    28
open ATP_Problem
blanchet@44428
    29
open ATP_Proof
blanchet@47148
    30
open ATP_Problem_Generate
blanchet@44428
    31
open ATP_Systems
blanchet@49232
    32
open ATP_Util
blanchet@43473
    33
blanchet@49228
    34
fun stringN_of_int 0 _ = ""
blanchet@49228
    35
  | stringN_of_int k n =
blanchet@49228
    36
    stringN_of_int (k - 1) (n div 10) ^ string_of_int (n mod 10)
blanchet@49228
    37
blanchet@49228
    38
fun escape_meta_char c =
blanchet@49228
    39
  if Char.isAlphaNum c orelse c = #"_" orelse c = #"." orelse c = #"(" orelse
blanchet@49248
    40
     c = #")" orelse c = #"," then
blanchet@49228
    41
    String.str c
blanchet@49232
    42
  else if c = #"'" then
blanchet@49232
    43
    "~"
blanchet@49228
    44
  else
blanchet@49228
    45
    (* fixed width, in case more digits follow *)
blanchet@49232
    46
    "\\" ^ stringN_of_int 3 (Char.ord c)
blanchet@49228
    47
val escape_meta = String.translate escape_meta_char
blanchet@49228
    48
blanchet@49240
    49
val thy_prefix = "y_"
blanchet@49240
    50
blanchet@49228
    51
val fact_name_of = escape_meta
blanchet@49240
    52
val thy_name_of = prefix thy_prefix o escape_meta
blanchet@49231
    53
val const_name_of = prefix const_prefix o escape_meta
blanchet@49231
    54
val type_name_of = prefix type_const_prefix o escape_meta
blanchet@49232
    55
val class_name_of = prefix class_prefix o escape_meta
blanchet@49231
    56
blanchet@49231
    57
val thy_name_of_thm = theory_of_thm #> Context.theory_name
blanchet@49231
    58
blanchet@49231
    59
fun has_thy thy th = (Context.theory_name thy = thy_name_of_thm th)
blanchet@43473
    60
blanchet@44329
    61
(* FIXME: Similar yet different code in "mirabelle.ML". The code here has a few
blanchet@44329
    62
   fixes that seem to be missing over there; or maybe the two code portions are
blanchet@44329
    63
   not doing the same? *)
blanchet@49228
    64
fun fold_body_thms thm_name f =
blanchet@43473
    65
  let
blanchet@44329
    66
    fun app n (PBody {thms, ...}) =
blanchet@44350
    67
      thms |> fold (fn (_, (name, prop, body)) => fn x =>
blanchet@44350
    68
        let
blanchet@44350
    69
          val body' = Future.join body
blanchet@44350
    70
          val n' =
blanchet@44350
    71
            n + (if name = "" orelse
blanchet@44369
    72
                    (* uncommon case where the proved theorem occurs twice
blanchet@44369
    73
                       (e.g., "Transitive_Closure.trancl_into_trancl") *)
blanchet@49228
    74
                    (n = 1 andalso name = thm_name) then
blanchet@44350
    75
                   0
blanchet@44350
    76
                 else
blanchet@44350
    77
                   1)
blanchet@44350
    78
          val x' = x |> n' <= 1 ? app n' body'
blanchet@44350
    79
        in (x' |> n = 1 ? f (name, prop, body')) end)
blanchet@44350
    80
  in fold (app 0) end
blanchet@43473
    81
blanchet@49228
    82
fun theorems_mentioned_in_proof_term all_names th =
blanchet@43473
    83
  let
blanchet@49228
    84
    val is_name_ok =
blanchet@49228
    85
      case all_names of
blanchet@49228
    86
        SOME names => member (op =) names
blanchet@49228
    87
      | NONE => (fn s => s <> "" andalso not (String.isPrefix "Pure." s))
blanchet@49228
    88
    fun collect (s, _, _) = is_name_ok s ? insert (op =) s
blanchet@43473
    89
    val names =
blanchet@49228
    90
      [] |> fold_body_thms (Thm.get_name_hint th) collect [Thm.proof_body_of th]
blanchet@44350
    91
         |> map fact_name_of
blanchet@43473
    92
  in names end
blanchet@43473
    93
blanchet@49243
    94
fun interesting_terms_types_and_classes term_max_depth type_max_depth t =
blanchet@49231
    95
  let
blanchet@49241
    96
    val bad_types = [@{type_name prop}, @{type_name bool}, @{type_name fun}]
blanchet@49241
    97
    val bad_consts = atp_widely_irrelevant_consts
blanchet@49232
    98
    val add_classes =
blanchet@49232
    99
      subtract (op =) @{sort type} #> map class_name_of #> union (op =)
blanchet@49243
   100
    fun do_add_type (Type (s, Ts)) =
blanchet@49241
   101
        (not (member (op =) bad_types s) ? insert (op =) (type_name_of s))
blanchet@49243
   102
        #> fold do_add_type Ts
blanchet@49243
   103
      | do_add_type (TFree (_, S)) = add_classes S
blanchet@49243
   104
      | do_add_type (TVar (_, S)) = add_classes S
blanchet@49243
   105
    fun add_type T = type_max_depth >= 0 ? do_add_type T
blanchet@49241
   106
    fun mk_app s args =
blanchet@49248
   107
      if member (op <>) args "" then s ^ "(" ^ space_implode "," args ^ ")"
blanchet@49241
   108
      else s
blanchet@49248
   109
    fun patternify ~1 _ = ""
blanchet@49241
   110
      | patternify depth t =
blanchet@49241
   111
        case strip_comb t of
blanchet@49241
   112
          (Const (s, _), args) =>
blanchet@49241
   113
          mk_app (const_name_of s) (map (patternify (depth - 1)) args)
blanchet@49248
   114
        | _ => ""
blanchet@49243
   115
    fun add_term_patterns ~1 _ = I
blanchet@49243
   116
      | add_term_patterns depth t =
blanchet@49241
   117
        insert (op =) (patternify depth t)
blanchet@49241
   118
        #> add_term_patterns (depth - 1) t
blanchet@49241
   119
    val add_term = add_term_patterns term_max_depth
blanchet@49241
   120
    fun add_patterns t =
blanchet@49241
   121
      let val (head, args) = strip_comb t in
blanchet@49241
   122
        (case head of
blanchet@49241
   123
           Const (s, T) =>
blanchet@49241
   124
           not (member (op =) bad_consts s) ? (add_term t #> add_type T)
blanchet@49241
   125
         | Free (_, T) => add_type T
blanchet@49241
   126
         | Var (_, T) => add_type T
blanchet@49241
   127
         | Abs (_, T, body) => add_type T #> add_patterns body
blanchet@49241
   128
         | _ => I)
blanchet@49241
   129
        #> fold add_patterns args
blanchet@49241
   130
      end
blanchet@49241
   131
  in [] |> add_patterns t |> sort string_ord end
blanchet@49231
   132
blanchet@49244
   133
fun is_likely_tautology th =
blanchet@49248
   134
  null (interesting_terms_types_and_classes 0 ~1 (prop_of th)) andalso
blanchet@49244
   135
  not (Thm.eq_thm_prop (@{thm ext}, th))
blanchet@49244
   136
blanchet@49244
   137
fun is_too_meta thy th =
blanchet@49244
   138
  fastype_of (Object_Logic.atomize_term thy (prop_of th)) <> @{typ bool}
blanchet@49244
   139
blanchet@49244
   140
fun facts_of thy =
blanchet@49244
   141
  let val ctxt = Proof_Context.init_global thy in
blanchet@49244
   142
    Sledgehammer_Filter.all_facts ctxt false Symtab.empty true [] []
blanchet@49244
   143
        (Sledgehammer_Filter.clasimpset_rule_table_of ctxt)
blanchet@49244
   144
    |> filter_out ((is_likely_tautology orf is_too_meta thy) o snd)
blanchet@49244
   145
    |> rev
blanchet@49244
   146
  end
blanchet@49244
   147
blanchet@49228
   148
fun theory_ord p =
blanchet@49228
   149
  if Theory.eq_thy p then EQUAL
blanchet@49228
   150
  else if Theory.subthy p then LESS
blanchet@49228
   151
  else if Theory.subthy (swap p) then GREATER
blanchet@49228
   152
  else EQUAL
blanchet@49228
   153
blanchet@49231
   154
val thm_ord = theory_ord o pairself theory_of_thm
blanchet@49231
   155
blanchet@49231
   156
fun parent_thms thy_ths thy =
blanchet@49231
   157
  Theory.parents_of thy
blanchet@49240
   158
  |> map Context.theory_name
blanchet@49231
   159
  |> map_filter (AList.lookup (op =) thy_ths)
blanchet@49231
   160
  |> map List.last
blanchet@49231
   161
  |> map (fact_name_of o Thm.get_name_hint)
blanchet@49231
   162
blanchet@49231
   163
val thms_by_thy =
blanchet@49231
   164
  map (snd #> `thy_name_of_thm)
blanchet@49231
   165
  #> AList.group (op =)
blanchet@49232
   166
  #> sort (int_ord
blanchet@49232
   167
           o pairself (length o Theory.ancestors_of o theory_of_thm o hd o snd))
blanchet@49231
   168
  #> map (apsnd (sort thm_ord))
blanchet@49231
   169
blanchet@49231
   170
fun generate_mash_accessibility_file_for_theory thy include_thy file_name =
blanchet@49228
   171
  let
blanchet@49228
   172
    val path = file_name |> Path.explode
blanchet@49228
   173
    val _ = File.write path ""
blanchet@49229
   174
    fun do_thm th prevs =
blanchet@49228
   175
      let
blanchet@49229
   176
        val s = th ^ ": " ^ space_implode " " prevs ^ "\n"
blanchet@49228
   177
        val _ = File.append path s
blanchet@49229
   178
      in [th] end
blanchet@49229
   179
    val thy_ths =
blanchet@49229
   180
      facts_of thy
blanchet@49231
   181
      |> not include_thy ? filter_out (has_thy thy o snd)
blanchet@49231
   182
      |> thms_by_thy
blanchet@49229
   183
    fun do_thy ths =
blanchet@49228
   184
      let
blanchet@49228
   185
        val thy = theory_of_thm (hd ths)
blanchet@49231
   186
        val parents = parent_thms thy_ths thy
blanchet@49229
   187
        val ths = ths |> map (fact_name_of o Thm.get_name_hint)
blanchet@49229
   188
        val _ = fold do_thm ths parents
blanchet@49228
   189
      in () end
blanchet@49229
   190
    val _ = List.app (do_thy o snd) thy_ths
blanchet@49228
   191
  in () end
blanchet@49228
   192
blanchet@49232
   193
fun has_bool @{typ bool} = true
blanchet@49232
   194
  | has_bool (Type (_, Ts)) = exists has_bool Ts
blanchet@49232
   195
  | has_bool _ = false
blanchet@49232
   196
blanchet@49232
   197
fun has_fun (Type (@{type_name fun}, _)) = true
blanchet@49232
   198
  | has_fun (Type (_, Ts)) = exists has_fun Ts
blanchet@49232
   199
  | has_fun _ = false
blanchet@49232
   200
blanchet@49232
   201
val is_conn = member (op =)
blanchet@49232
   202
  [@{const_name Trueprop}, @{const_name HOL.conj}, @{const_name HOL.disj},
blanchet@49232
   203
   @{const_name HOL.implies}, @{const_name Not},
blanchet@49232
   204
   @{const_name All}, @{const_name Ex}, @{const_name Ball}, @{const_name Bex},
blanchet@49232
   205
   @{const_name HOL.eq}]
blanchet@49232
   206
blanchet@49232
   207
val has_bool_arg_const =
blanchet@49232
   208
  exists_Const (fn (c, T) =>
blanchet@49232
   209
                   not (is_conn c) andalso exists has_bool (binder_types T))
blanchet@49232
   210
blanchet@49232
   211
fun higher_inst_const thy (c, T) =
blanchet@49232
   212
  case binder_types T of
blanchet@49232
   213
    [] => false
blanchet@49232
   214
  | Ts => length (binder_types (Sign.the_const_type thy c)) <> length Ts
blanchet@49232
   215
blanchet@49232
   216
val binders = [@{const_name All}, @{const_name Ex}]
blanchet@49232
   217
blanchet@49232
   218
fun is_fo_term thy t =
blanchet@49232
   219
  let
blanchet@49232
   220
    val t =
blanchet@49232
   221
      t |> Envir.beta_eta_contract
blanchet@49232
   222
        |> transform_elim_prop
blanchet@49232
   223
        |> Object_Logic.atomize_term thy
blanchet@49232
   224
  in
blanchet@49232
   225
    Term.is_first_order binders t andalso
blanchet@49232
   226
    not (exists_subterm (fn Var (_, T) => has_bool T orelse has_fun T
blanchet@49232
   227
                          | _ => false) t orelse
blanchet@49232
   228
         has_bool_arg_const t orelse exists_Const (higher_inst_const thy) t)
blanchet@49232
   229
  end
blanchet@49232
   230
blanchet@49243
   231
fun is_exists (s, _) = (s = @{const_name Ex} orelse s = @{const_name Ex1})
blanchet@49232
   232
blanchet@49241
   233
val max_depth = 1
blanchet@49241
   234
blanchet@49231
   235
fun features_of thy (status, th) =
blanchet@49243
   236
  let val t = Thm.prop_of th in
blanchet@49240
   237
    thy_name_of (thy_name_of_thm th) ::
blanchet@49243
   238
    interesting_terms_types_and_classes max_depth max_depth t
blanchet@49243
   239
    |> not (has_no_lambdas t) ? cons "lambdas"
blanchet@49243
   240
    |> exists_Const is_exists t ? cons "skolems"
blanchet@49243
   241
    |> not (is_fo_term thy t) ? cons "ho"
blanchet@49231
   242
    |> (case status of
blanchet@49231
   243
          General => I
blanchet@49231
   244
        | Induction => cons "induction"
blanchet@49231
   245
        | Intro => cons "intro"
blanchet@49231
   246
        | Inductive => cons "inductive"
blanchet@49231
   247
        | Elim => cons "elim"
blanchet@49231
   248
        | Simp => cons "simp"
blanchet@49231
   249
        | Def => cons "def")
blanchet@49231
   250
  end
blanchet@49231
   251
blanchet@49231
   252
fun generate_mash_feature_file_for_theory thy include_thy file_name =
blanchet@43473
   253
  let
blanchet@43473
   254
    val path = file_name |> Path.explode
blanchet@43473
   255
    val _ = File.write path ""
blanchet@49231
   256
    val facts = facts_of thy |> not include_thy ? filter_out (has_thy thy o snd)
blanchet@49230
   257
    fun do_fact ((_, (_, status)), th) =
blanchet@43473
   258
      let
blanchet@49228
   259
        val name = Thm.get_name_hint th
blanchet@49231
   260
        val feats = features_of thy (status, th)
blanchet@49231
   261
        val s = fact_name_of name ^ ": " ^ space_implode " " feats ^ "\n"
blanchet@43473
   262
      in File.append path s end
blanchet@49230
   263
    val _ = List.app do_fact facts
blanchet@49228
   264
  in () end
blanchet@49228
   265
blanchet@49231
   266
val dependencies_of = theorems_mentioned_in_proof_term o SOME
blanchet@49231
   267
blanchet@49231
   268
fun generate_mash_dependency_file_for_theory thy include_thy file_name =
blanchet@49228
   269
  let
blanchet@49228
   270
    val path = file_name |> Path.explode
blanchet@49228
   271
    val _ = File.write path ""
blanchet@49231
   272
    val ths =
blanchet@49231
   273
      facts_of thy |> not include_thy ? filter_out (has_thy thy o snd)
blanchet@49231
   274
                   |> map snd
blanchet@49244
   275
    val all_names = ths |> map Thm.get_name_hint
blanchet@49228
   276
    fun do_thm th =
blanchet@49228
   277
      let
blanchet@49228
   278
        val name = Thm.get_name_hint th
blanchet@49231
   279
        val deps = dependencies_of all_names th
blanchet@49231
   280
        val s = fact_name_of name ^ ": " ^ space_implode " " deps ^ "\n"
blanchet@49228
   281
      in File.append path s end
blanchet@49228
   282
    val _ = List.app do_thm ths
blanchet@43473
   283
  in () end
blanchet@43473
   284
blanchet@49231
   285
fun generate_mash_problem_file_for_theory thy file_name =
blanchet@49231
   286
  let
blanchet@49231
   287
    val path = file_name |> Path.explode
blanchet@49231
   288
    val _ = File.write path ""
blanchet@49231
   289
    val facts = facts_of thy
blanchet@49231
   290
    val (new_facts, old_facts) =
blanchet@49231
   291
      facts |> List.partition (has_thy thy o snd)
blanchet@49231
   292
            |>> sort (thm_ord o pairself snd)
blanchet@49232
   293
    val ths = facts |> map snd
blanchet@49244
   294
    val all_names = ths |> map Thm.get_name_hint
blanchet@49231
   295
    fun do_fact ((_, (_, status)), th) prevs =
blanchet@49231
   296
      let
blanchet@49231
   297
        val name = Thm.get_name_hint th
blanchet@49231
   298
        val feats = features_of thy (status, th)
blanchet@49231
   299
        val deps = dependencies_of all_names th
blanchet@49248
   300
        val kind = Thm.legacy_get_kind th
blanchet@49248
   301
        val name = fact_name_of name
blanchet@49248
   302
        val core =
blanchet@49248
   303
          name ^ ": " ^ space_implode " " prevs ^ "; " ^ space_implode " " feats
blanchet@49248
   304
        val query = if kind <> "" then "? " ^ core ^ "\n" else ""
blanchet@49248
   305
        val update = "! " ^ core ^ "; " ^ space_implode " " deps ^ "\n"
blanchet@49248
   306
        val _ = File.append path (query ^ update)
blanchet@49248
   307
      in [name] end
blanchet@49231
   308
    val thy_ths = old_facts |> thms_by_thy
blanchet@49231
   309
    val parents = parent_thms thy_ths thy
blanchet@49231
   310
    val _ = fold do_fact new_facts parents
blanchet@49231
   311
  in () end
blanchet@49231
   312
blanchet@43473
   313
fun inference_term [] = NONE
blanchet@43473
   314
  | inference_term ss =
blanchet@49147
   315
    ATerm (("inference", []),
blanchet@49147
   316
           [ATerm (("isabelle", []), []),
blanchet@49147
   317
            ATerm ((tptp_empty_list, []), []),
blanchet@49147
   318
            ATerm ((tptp_empty_list, []),
blanchet@49147
   319
            map (fn s => ATerm ((s, []), [])) ss)])
blanchet@43473
   320
    |> SOME
blanchet@43473
   321
fun inference infers ident =
blanchet@43473
   322
  these (AList.lookup (op =) infers ident) |> inference_term
blanchet@43473
   323
fun add_inferences_to_problem_line infers
blanchet@47234
   324
                                   (Formula (ident, Axiom, phi, NONE, tms)) =
blanchet@47234
   325
    Formula (ident, Lemma, phi, inference infers ident, tms)
blanchet@43473
   326
  | add_inferences_to_problem_line _ line = line
blanchet@44867
   327
fun add_inferences_to_problem infers =
blanchet@44867
   328
  map (apsnd (map (add_inferences_to_problem_line infers)))
blanchet@43473
   329
blanchet@49157
   330
fun ident_of_problem_line (Class_Decl (ident, _, _)) = ident
blanchet@49157
   331
  | ident_of_problem_line (Type_Decl (ident, _, _)) = ident
blanchet@49152
   332
  | ident_of_problem_line (Sym_Decl (ident, _, _)) = ident
blanchet@49157
   333
  | ident_of_problem_line (Class_Memb (ident, _, _, _)) = ident
blanchet@44350
   334
  | ident_of_problem_line (Formula (ident, _, _, _, _)) = ident
blanchet@44350
   335
blanchet@46176
   336
fun run_some_atp ctxt format problem =
blanchet@44428
   337
  let
blanchet@44428
   338
    val thy = Proof_Context.theory_of ctxt
blanchet@47925
   339
    val prob_file = File.tmp_path (Path.explode "prob")
blanchet@49146
   340
    val atp = case format of DFG _ => spassN | _ => eN
blanchet@48469
   341
    val {exec, arguments, proof_delims, known_failures, ...} =
blanchet@48469
   342
      get_atp thy atp ()
blanchet@47909
   343
    val ord = effective_term_order ctxt atp
blanchet@47909
   344
    val _ = problem |> lines_for_atp_problem format ord (K [])
blanchet@47270
   345
                    |> File.write_list prob_file
blanchet@49228
   346
    val path = getenv (List.last (fst exec)) ^ "/" ^ snd exec
blanchet@44428
   347
    val command =
blanchet@49228
   348
      File.shell_path (Path.explode path) ^
blanchet@47909
   349
      " " ^ arguments ctxt false "" (seconds 1.0) (ord, K [], K []) ^ " " ^
blanchet@44428
   350
      File.shell_path prob_file
blanchet@44428
   351
  in
wenzelm@44721
   352
    TimeLimit.timeLimit (seconds 0.3) Isabelle_System.bash_output command
blanchet@44428
   353
    |> fst
blanchet@47909
   354
    |> extract_tstplike_proof_and_outcome false true proof_delims known_failures
blanchet@44428
   355
    |> snd
blanchet@44428
   356
  end
blanchet@44428
   357
  handle TimeLimit.TimeOut => SOME TimedOut
blanchet@44428
   358
blanchet@49232
   359
val tautology_prefixes =
blanchet@44431
   360
  [@{theory HOL}, @{theory Meson}, @{theory ATP}, @{theory Metis}]
blanchet@44428
   361
  |> map (fact_name_of o Context.theory_name)
blanchet@44428
   362
blanchet@46176
   363
fun is_problem_line_tautology ctxt format (Formula (ident, _, phi, _, _)) =
blanchet@44428
   364
    exists (fn prefix => String.isPrefix prefix ident)
blanchet@49232
   365
           tautology_prefixes andalso
blanchet@46176
   366
    is_none (run_some_atp ctxt format
blanchet@47234
   367
                 [(factsN, [Formula (ident, Conjecture, phi, NONE, [])])])
blanchet@46176
   368
  | is_problem_line_tautology _ _ _ = false
blanchet@44428
   369
blanchet@44370
   370
fun order_facts ord = sort (ord o pairself ident_of_problem_line)
blanchet@44370
   371
fun order_problem_facts _ [] = []
blanchet@44370
   372
  | order_problem_facts ord ((heading, lines) :: problem) =
blanchet@44370
   373
    if heading = factsN then (heading, order_facts ord lines) :: problem
blanchet@44370
   374
    else (heading, lines) :: order_problem_facts ord problem
blanchet@44370
   375
blanchet@46176
   376
(* A fairly random selection of types used for monomorphizing. *)
blanchet@46176
   377
val ground_types =
blanchet@46176
   378
  [@{typ nat}, HOLogic.intT, HOLogic.realT, @{typ "nat => bool"}, @{typ bool},
blanchet@46176
   379
   @{typ unit}]
blanchet@46176
   380
blanchet@46176
   381
fun ground_type_for_tvar _ [] tvar =
blanchet@46176
   382
    raise TYPE ("ground_type_for_sorts", [TVar tvar], [])
blanchet@46176
   383
  | ground_type_for_tvar thy (T :: Ts) tvar =
blanchet@46176
   384
    if can (Sign.typ_match thy (TVar tvar, T)) Vartab.empty then T
blanchet@46176
   385
    else ground_type_for_tvar thy Ts tvar
blanchet@46176
   386
blanchet@46176
   387
fun monomorphize_term ctxt t =
blanchet@46176
   388
  let val thy = Proof_Context.theory_of ctxt in
blanchet@46176
   389
    t |> map_types (map_type_tvar (ground_type_for_tvar thy ground_types))
blanchet@46176
   390
    handle TYPE _ => @{prop True}
blanchet@46176
   391
  end
blanchet@46176
   392
blanchet@46176
   393
fun generate_tptp_inference_file_for_theory ctxt thy format type_enc file_name =
blanchet@43473
   394
  let
blanchet@47129
   395
    val type_enc = type_enc |> type_enc_from_string Strict
blanchet@46176
   396
                            |> adjust_type_enc format
blanchet@49146
   397
    val mono = not (is_type_enc_polymorphic type_enc)
blanchet@43473
   398
    val path = file_name |> Path.explode
blanchet@43473
   399
    val _ = File.write path ""
blanchet@44438
   400
    val facts = facts_of thy
blanchet@46422
   401
    val atp_problem =
blanchet@44438
   402
      facts
blanchet@46176
   403
      |> map (fn ((_, loc), th) =>
blanchet@46176
   404
                 ((Thm.get_name_hint th, loc),
blanchet@46176
   405
                   th |> prop_of |> mono ? monomorphize_term ctxt))
blanchet@48961
   406
      |> prepare_atp_problem ctxt format Axiom type_enc Exporter combsN false
blanchet@48961
   407
                             false true [] @{prop False}
blanchet@46422
   408
      |> #1
blanchet@44428
   409
    val atp_problem =
blanchet@44428
   410
      atp_problem
blanchet@46176
   411
      |> map (apsnd (filter_out (is_problem_line_tautology ctxt format)))
blanchet@49232
   412
    val ths = facts |> map snd
blanchet@49244
   413
    val all_names = ths |> map Thm.get_name_hint
blanchet@43473
   414
    val infers =
blanchet@44438
   415
      facts |> map (fn (_, th) =>
blanchet@44438
   416
                       (fact_name_of (Thm.get_name_hint th),
blanchet@44438
   417
                        theorems_mentioned_in_proof_term (SOME all_names) th))
blanchet@44350
   418
    val all_atp_problem_names =
blanchet@44350
   419
      atp_problem |> maps (map ident_of_problem_line o snd)
blanchet@43473
   420
    val infers =
blanchet@44370
   421
      infers |> filter (member (op =) all_atp_problem_names o fst)
blanchet@44370
   422
             |> map (apsnd (filter (member (op =) all_atp_problem_names)))
blanchet@44370
   423
    val ordered_names =
blanchet@44370
   424
      String_Graph.empty
blanchet@44370
   425
      |> fold (String_Graph.new_node o rpair ()) all_atp_problem_names
blanchet@44370
   426
      |> fold (fn (to, froms) =>
blanchet@44428
   427
                  fold (fn from => String_Graph.add_edge (from, to)) froms)
blanchet@44428
   428
              infers
blanchet@44370
   429
      |> String_Graph.topological_order
blanchet@44370
   430
    val order_tab =
blanchet@44370
   431
      Symtab.empty
blanchet@44370
   432
      |> fold (Symtab.insert (op =))
blanchet@44370
   433
              (ordered_names ~~ (1 upto length ordered_names))
blanchet@44370
   434
    val name_ord = int_ord o pairself (the o Symtab.lookup order_tab)
blanchet@44370
   435
    val atp_problem =
blanchet@46176
   436
      atp_problem
blanchet@49146
   437
      |> (case format of DFG _ => I | _ => add_inferences_to_problem infers)
blanchet@46176
   438
      |> order_problem_facts name_ord
blanchet@47909
   439
    val ord = effective_term_order ctxt eN (* dummy *)
blanchet@47909
   440
    val ss = lines_for_atp_problem format ord (K []) atp_problem
blanchet@43473
   441
    val _ = app (File.append path) ss
blanchet@43473
   442
  in () end
blanchet@43473
   443
blanchet@43473
   444
end;