src/HOL/TPTP/atp_theory_export.ML
author blanchet
Mon, 09 Jul 2012 23:23:12 +0200
changeset 48213 d20add034f64
parent 48142 efaff8206967
child 48214 36348e75af66
permissions -rw-r--r--
first go at generating files for MaSh (machine-learning Sledgehammer)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
46321
484dc68c8c89 renamed theory exporter
blanchet
parents: 46320
diff changeset
     1
(*  Title:      HOL/TPTP/atp_theory_export.ML
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
     3
    Copyright   2011
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
     4
48213
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
     5
Export Isabelle theories as MaSh (Machine-learning for Sledgehammer) or as
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
     6
first-order TPTP inferences.
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
     7
*)
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
     8
46321
484dc68c8c89 renamed theory exporter
blanchet
parents: 46320
diff changeset
     9
signature ATP_THEORY_EXPORT =
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    10
sig
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
    11
  type atp_format = ATP_Problem.atp_format
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
    12
43494
13eefebbc4cb make sure that enough type information is generated -- because the exported "lemma"s are also used as "conjecture", we can't optimize type information based on polarity
blanchet
parents: 43479
diff changeset
    13
  val theorems_mentioned_in_proof_term :
13eefebbc4cb make sure that enough type information is generated -- because the exported "lemma"s are also used as "conjecture", we can't optimize type information based on polarity
blanchet
parents: 43479
diff changeset
    14
    string list option -> thm -> string list
48213
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    15
  val generate_mash_accessibility_file_for_theory : theory -> string -> unit
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    16
  val generate_mash_feature_file_for_theory :
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    17
    Proof.context -> theory -> string -> unit
48213
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    18
  val generate_mash_dependency_file_for_theory : theory -> string -> unit
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    19
  val generate_tptp_inference_file_for_theory :
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
    20
    Proof.context -> theory -> atp_format -> string -> string -> unit
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    21
end;
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    22
46321
484dc68c8c89 renamed theory exporter
blanchet
parents: 46320
diff changeset
    23
structure ATP_Theory_Export : ATP_THEORY_EXPORT =
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    24
struct
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    25
43479
5af1abc13c1f only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
blanchet
parents: 43468
diff changeset
    26
open ATP_Problem
43566
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
    27
open ATP_Proof
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 46301
diff changeset
    28
open ATP_Problem_Generate
43566
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
    29
open ATP_Systems
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    30
48213
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    31
val thy_prefix = "$"
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    32
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    33
fun stringN_of_int 0 _ = ""
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    34
  | stringN_of_int k n =
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    35
    stringN_of_int (k - 1) (n div 10) ^ string_of_int (n mod 10)
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    36
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    37
fun escape_meta_char c =
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    38
  if Char.isAlphaNum c orelse c = #"_" orelse c = #"." orelse c = #"(" orelse
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    39
     c = #")" then
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    40
    String.str c
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    41
  else
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    42
    (* fixed width, in case more digits follow *)
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    43
    "_" ^ stringN_of_int 3 (Char.ord c)
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    44
val escape_meta = String.translate escape_meta_char
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    45
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    46
val fact_name_of = escape_meta
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    47
val thy_name_of = prefix thy_prefix o escape_meta
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    48
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    49
fun facts_of thy =
46734
6256e064f8fa speed up Sledgehammer's clasimpset lookup a bit
blanchet
parents: 46667
diff changeset
    50
  let val ctxt = Proof_Context.init_global thy in
6256e064f8fa speed up Sledgehammer's clasimpset lookup a bit
blanchet
parents: 46667
diff changeset
    51
    Sledgehammer_Filter.all_facts ctxt false
6256e064f8fa speed up Sledgehammer's clasimpset lookup a bit
blanchet
parents: 46667
diff changeset
    52
        Symtab.empty true [] []
6256e064f8fa speed up Sledgehammer's clasimpset lookup a bit
blanchet
parents: 46667
diff changeset
    53
        (Sledgehammer_Filter.clasimpset_rule_table_of ctxt)
6256e064f8fa speed up Sledgehammer's clasimpset lookup a bit
blanchet
parents: 46667
diff changeset
    54
    |> filter (curry (op =) @{typ bool} o fastype_of
6256e064f8fa speed up Sledgehammer's clasimpset lookup a bit
blanchet
parents: 46667
diff changeset
    55
               o Object_Logic.atomize_term thy o prop_of o snd)
6256e064f8fa speed up Sledgehammer's clasimpset lookup a bit
blanchet
parents: 46667
diff changeset
    56
  end
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    57
43468
c768f7adb711 fixed silly ATP exporter bug: if the proof of lemma A relies on B and C, and the proof of B relies on C, return {B, C}, not {B}, as the set of dependencies
blanchet
parents: 43400
diff changeset
    58
(* FIXME: Similar yet different code in "mirabelle.ML". The code here has a few
c768f7adb711 fixed silly ATP exporter bug: if the proof of lemma A relies on B and C, and the proof of B relies on C, return {B, C}, not {B}, as the set of dependencies
blanchet
parents: 43400
diff changeset
    59
   fixes that seem to be missing over there; or maybe the two code portions are
c768f7adb711 fixed silly ATP exporter bug: if the proof of lemma A relies on B and C, and the proof of B relies on C, return {B, C}, not {B}, as the set of dependencies
blanchet
parents: 43400
diff changeset
    60
   not doing the same? *)
48213
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    61
fun fold_body_thms thm_name f =
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    62
  let
43468
c768f7adb711 fixed silly ATP exporter bug: if the proof of lemma A relies on B and C, and the proof of B relies on C, return {B, C}, not {B}, as the set of dependencies
blanchet
parents: 43400
diff changeset
    63
    fun app n (PBody {thms, ...}) =
43479
5af1abc13c1f only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
blanchet
parents: 43468
diff changeset
    64
      thms |> fold (fn (_, (name, prop, body)) => fn x =>
5af1abc13c1f only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
blanchet
parents: 43468
diff changeset
    65
        let
5af1abc13c1f only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
blanchet
parents: 43468
diff changeset
    66
          val body' = Future.join body
5af1abc13c1f only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
blanchet
parents: 43468
diff changeset
    67
          val n' =
5af1abc13c1f only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
blanchet
parents: 43468
diff changeset
    68
            n + (if name = "" orelse
43498
75caf7e4302e peel off two or more layers in exceptional cases where the proof term refers to the proved theorems twice with the same name (e.g., "Transitive_Closure.trancl_into_trancl")
blanchet
parents: 43494
diff changeset
    69
                    (* uncommon case where the proved theorem occurs twice
75caf7e4302e peel off two or more layers in exceptional cases where the proof term refers to the proved theorems twice with the same name (e.g., "Transitive_Closure.trancl_into_trancl")
blanchet
parents: 43494
diff changeset
    70
                       (e.g., "Transitive_Closure.trancl_into_trancl") *)
48213
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    71
                    (n = 1 andalso name = thm_name) then
43479
5af1abc13c1f only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
blanchet
parents: 43468
diff changeset
    72
                   0
5af1abc13c1f only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
blanchet
parents: 43468
diff changeset
    73
                 else
5af1abc13c1f only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
blanchet
parents: 43468
diff changeset
    74
                   1)
5af1abc13c1f only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
blanchet
parents: 43468
diff changeset
    75
          val x' = x |> n' <= 1 ? app n' body'
5af1abc13c1f only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
blanchet
parents: 43468
diff changeset
    76
        in (x' |> n = 1 ? f (name, prop, body')) end)
5af1abc13c1f only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
blanchet
parents: 43468
diff changeset
    77
  in fold (app 0) end
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    78
48213
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    79
fun theorems_mentioned_in_proof_term all_names th =
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    80
  let
48213
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    81
    val is_name_ok =
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    82
      case all_names of
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    83
        SOME names => member (op =) names
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    84
      | NONE => (fn s => s <> "" andalso not (String.isPrefix "Pure." s))
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    85
    fun collect (s, _, _) = is_name_ok s ? insert (op =) s
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    86
    val names =
48213
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    87
      [] |> fold_body_thms (Thm.get_name_hint th) collect [Thm.proof_body_of th]
43479
5af1abc13c1f only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
blanchet
parents: 43468
diff changeset
    88
         |> map fact_name_of
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    89
  in names end
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    90
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    91
fun interesting_const_names ctxt =
43276
91bf67e0e755 modernized Proof_Context;
wenzelm
parents: 43259
diff changeset
    92
  let val thy = Proof_Context.theory_of ctxt in
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    93
    Sledgehammer_Filter.const_names_in_fact thy
43566
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
    94
        (Sledgehammer_Provers.is_built_in_const_for_prover ctxt eN)
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    95
  end
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    96
48213
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    97
fun theory_ord p =
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    98
  if Theory.eq_thy p then EQUAL
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    99
  else if Theory.subthy p then LESS
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   100
  else if Theory.subthy (swap p) then GREATER
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   101
  else EQUAL
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   102
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   103
fun generate_mash_accessibility_file_for_theory thy file_name =
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   104
  let
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   105
    val path = file_name |> Path.explode
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   106
    val _ = File.write path ""
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   107
    fun do_thm parents th seen =
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   108
      let
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   109
        val s = th ^ ": " ^ space_implode " " (parents @ seen) ^ "\n"
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   110
        val _ = File.append path s
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   111
      in seen @ [th] end
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   112
    fun do_thy (name, ths) =
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   113
      let
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   114
        val ths = sort (theory_ord o pairself theory_of_thm) ths
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   115
        val thy = theory_of_thm (hd ths)
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   116
        val parents =
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   117
          if Context.theory_name thy = Context.theory_name @{theory HOL} then []
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   118
          else map (thy_name_of o Context.theory_name) (Theory.parents_of thy)
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   119
        val ths = map (fact_name_of o Thm.get_name_hint) ths
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   120
        val s =
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   121
          thy_name_of name ^ ": " ^ space_implode " " (parents @ ths) ^ " \n"
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   122
        val _ = File.append path s
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   123
        val _ = fold (do_thm parents) ths []
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   124
      in () end
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   125
    val thy_name_of_thm = theory_of_thm #> Context.theory_name
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   126
    val thy_ths =
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   127
      facts_of thy |> map (snd #> `thy_name_of_thm)
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   128
                   |> AList.group (op =)
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   129
                   |> sort (theory_ord o pairself (theory_of_thm o hd o snd))
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   130
    val _ = List.app do_thy thy_ths
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   131
  in () end
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   132
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   133
(* TODO: Add types, subterms, intro/dest/elim/simp status *)
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   134
fun generate_mash_feature_file_for_theory ctxt thy file_name =
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   135
  let
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   136
    val path = file_name |> Path.explode
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   137
    val _ = File.write path ""
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   138
    val axioms = Theory.all_axioms_of thy |> map fst
48213
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   139
    fun do_thm th =
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   140
      let
48213
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   141
        val name = Thm.get_name_hint th
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   142
        val features =
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   143
          (if member (op =) axioms name then ["axiom"] else []) @
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   144
          map (prefix const_prefix o escape_meta)
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   145
              (interesting_const_names ctxt (Thm.prop_of th))
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   146
          |> (fn [] => ["likely_tautology"] | features => features)
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   147
        val s = fact_name_of name ^ ": " ^ space_implode " " features ^ "\n"
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   148
      in File.append path s end
48213
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   149
    val ths = facts_of thy |> map snd
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   150
    val _ = List.app do_thm ths
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   151
  in () end
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   152
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   153
fun generate_mash_dependency_file_for_theory thy file_name =
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   154
  let
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   155
    val path = file_name |> Path.explode
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   156
    val _ = File.write path ""
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   157
    val ths = facts_of thy |> map snd
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   158
    val all_names = map Thm.get_name_hint ths
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   159
    fun do_thm th =
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   160
      let
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   161
        val name = Thm.get_name_hint th
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   162
        val ths = theorems_mentioned_in_proof_term (SOME all_names) th
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   163
        val s = fact_name_of name ^ ": " ^ space_implode " " ths ^ "\n"
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   164
      in File.append path s end
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   165
    val _ = List.app do_thm ths
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   166
  in () end
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   167
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   168
fun inference_term [] = NONE
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   169
  | inference_term ss =
48132
9aa0fad4e864 added type arguments to "ATerm" constructor -- but don't use them yet
blanchet
parents: 48131
diff changeset
   170
    ATerm (("inference", []),
9aa0fad4e864 added type arguments to "ATerm" constructor -- but don't use them yet
blanchet
parents: 48131
diff changeset
   171
           [ATerm (("isabelle", []), []),
9aa0fad4e864 added type arguments to "ATerm" constructor -- but don't use them yet
blanchet
parents: 48131
diff changeset
   172
            ATerm ((tptp_empty_list, []), []),
9aa0fad4e864 added type arguments to "ATerm" constructor -- but don't use them yet
blanchet
parents: 48131
diff changeset
   173
            ATerm ((tptp_empty_list, []),
9aa0fad4e864 added type arguments to "ATerm" constructor -- but don't use them yet
blanchet
parents: 48131
diff changeset
   174
            map (fn s => ATerm ((s, []), [])) ss)])
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   175
    |> SOME
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   176
fun inference infers ident =
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   177
  these (AList.lookup (op =) infers ident) |> inference_term
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   178
fun add_inferences_to_problem_line infers
46406
0e490b9e8422 extended SPASS/DFG output with ranks
blanchet
parents: 46365
diff changeset
   179
                                   (Formula (ident, Axiom, phi, NONE, tms)) =
0e490b9e8422 extended SPASS/DFG output with ranks
blanchet
parents: 46365
diff changeset
   180
    Formula (ident, Lemma, phi, inference infers ident, tms)
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   181
  | add_inferences_to_problem_line _ line = line
43996
4d1270ddf042 make SML/NJ happy
blanchet
parents: 43863
diff changeset
   182
fun add_inferences_to_problem infers =
4d1270ddf042 make SML/NJ happy
blanchet
parents: 43863
diff changeset
   183
  map (apsnd (map (add_inferences_to_problem_line infers)))
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   184
48142
efaff8206967 finished implementation of DFG type class output
blanchet
parents: 48140
diff changeset
   185
fun ident_of_problem_line (Class_Decl (ident, _, _)) = ident
efaff8206967 finished implementation of DFG type class output
blanchet
parents: 48140
diff changeset
   186
  | ident_of_problem_line (Type_Decl (ident, _, _)) = ident
48137
6f524f2066e3 cleanly distinguish between type declarations and symbol declarations
blanchet
parents: 48132
diff changeset
   187
  | ident_of_problem_line (Sym_Decl (ident, _, _)) = ident
48142
efaff8206967 finished implementation of DFG type class output
blanchet
parents: 48140
diff changeset
   188
  | ident_of_problem_line (Class_Memb (ident, _, _, _)) = ident
43479
5af1abc13c1f only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
blanchet
parents: 43468
diff changeset
   189
  | ident_of_problem_line (Formula (ident, _, _, _, _)) = ident
5af1abc13c1f only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
blanchet
parents: 43468
diff changeset
   190
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   191
fun run_some_atp ctxt format problem =
43566
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
   192
  let
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
   193
    val thy = Proof_Context.theory_of ctxt
47055
16e2633f3b4b made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents: 47038
diff changeset
   194
    val prob_file = File.tmp_path (Path.explode "prob")
48131
1016664b8feb started adding polymophic SPASS output
blanchet
parents: 48130
diff changeset
   195
    val atp = case format of DFG _ => spassN | _ => eN
47606
06dde48a1503 true delayed evaluation of "SPASS_VERSION" environment variable
blanchet
parents: 47055
diff changeset
   196
    val {exec, arguments, proof_delims, known_failures, ...} =
06dde48a1503 true delayed evaluation of "SPASS_VERSION" environment variable
blanchet
parents: 47055
diff changeset
   197
      get_atp thy atp ()
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 46734
diff changeset
   198
    val ord = effective_term_order ctxt atp
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 46734
diff changeset
   199
    val _ = problem |> lines_for_atp_problem format ord (K [])
46442
1e07620d724c added possibility of generating KBO weights to DFG problems
blanchet
parents: 46409
diff changeset
   200
                    |> File.write_list prob_file
48213
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   201
    val path = getenv (List.last (fst exec)) ^ "/" ^ snd exec
43566
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
   202
    val command =
48213
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
   203
      File.shell_path (Path.explode path) ^
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 46734
diff changeset
   204
      " " ^ arguments ctxt false "" (seconds 1.0) (ord, K [], K []) ^ " " ^
43566
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
   205
      File.shell_path prob_file
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
   206
  in
43850
7f2cbc713344 moved bash operations to Isabelle_System (cf. Scala version);
wenzelm
parents: 43828
diff changeset
   207
    TimeLimit.timeLimit (seconds 0.3) Isabelle_System.bash_output command
43566
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
   208
    |> fst
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 46734
diff changeset
   209
    |> extract_tstplike_proof_and_outcome false true proof_delims known_failures
43566
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
   210
    |> snd
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
   211
  end
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
   212
  handle TimeLimit.TimeOut => SOME TimedOut
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
   213
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
   214
val likely_tautology_prefixes =
43569
b342cd125533 removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
blanchet
parents: 43566
diff changeset
   215
  [@{theory HOL}, @{theory Meson}, @{theory ATP}, @{theory Metis}]
43566
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
   216
  |> map (fact_name_of o Context.theory_name)
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
   217
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   218
fun is_problem_line_tautology ctxt format (Formula (ident, _, phi, _, _)) =
43566
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
   219
    exists (fn prefix => String.isPrefix prefix ident)
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
   220
           likely_tautology_prefixes andalso
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   221
    is_none (run_some_atp ctxt format
46406
0e490b9e8422 extended SPASS/DFG output with ranks
blanchet
parents: 46365
diff changeset
   222
                 [(factsN, [Formula (ident, Conjecture, phi, NONE, [])])])
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   223
  | is_problem_line_tautology _ _ _ = false
43566
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
   224
43499
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   225
fun order_facts ord = sort (ord o pairself ident_of_problem_line)
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   226
fun order_problem_facts _ [] = []
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   227
  | order_problem_facts ord ((heading, lines) :: problem) =
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   228
    if heading = factsN then (heading, order_facts ord lines) :: problem
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   229
    else (heading, lines) :: order_problem_facts ord problem
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   230
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   231
(* A fairly random selection of types used for monomorphizing. *)
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   232
val ground_types =
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   233
  [@{typ nat}, HOLogic.intT, HOLogic.realT, @{typ "nat => bool"}, @{typ bool},
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   234
   @{typ unit}]
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   235
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   236
fun ground_type_for_tvar _ [] tvar =
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   237
    raise TYPE ("ground_type_for_sorts", [TVar tvar], [])
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   238
  | ground_type_for_tvar thy (T :: Ts) tvar =
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   239
    if can (Sign.typ_match thy (TVar tvar, T)) Vartab.empty then T
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   240
    else ground_type_for_tvar thy Ts tvar
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   241
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   242
fun monomorphize_term ctxt t =
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   243
  let val thy = Proof_Context.theory_of ctxt in
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   244
    t |> map_types (map_type_tvar (ground_type_for_tvar thy ground_types))
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   245
    handle TYPE _ => @{prop True}
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   246
  end
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   247
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   248
fun generate_tptp_inference_file_for_theory ctxt thy format type_enc file_name =
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   249
  let
46301
e2e52c7d25c9 renamed "sound" option to "strict"
blanchet
parents: 45551
diff changeset
   250
    val type_enc = type_enc |> type_enc_from_string Strict
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   251
                            |> adjust_type_enc format
48131
1016664b8feb started adding polymophic SPASS output
blanchet
parents: 48130
diff changeset
   252
    val mono = not (is_type_enc_polymorphic type_enc)
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   253
    val path = file_name |> Path.explode
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   254
    val _ = File.write path ""
43576
ebeda6275027 don't export any metastrange or other nonatomizable formulas, since these don't help proving normal things, they are somewhat broken in the ATP output, and they are atypical
blanchet
parents: 43572
diff changeset
   255
    val facts = facts_of thy
45551
a62c7a21f4ab removed needless baggage
blanchet
parents: 45508
diff changeset
   256
    val atp_problem =
43576
ebeda6275027 don't export any metastrange or other nonatomizable formulas, since these don't help proving normal things, they are somewhat broken in the ATP output, and they are atypical
blanchet
parents: 43572
diff changeset
   257
      facts
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   258
      |> map (fn ((_, loc), th) =>
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   259
                 ((Thm.get_name_hint th, loc),
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   260
                   th |> prop_of |> mono ? monomorphize_term ctxt))
47946
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 47912
diff changeset
   261
      |> prepare_atp_problem ctxt format Axiom type_enc Exporter combsN false
33afcfad3f8d add an experimental "aggressive" mode to Sledgehammer, to experiment with more complete translations of higher-order features without breaking "metis"
blanchet
parents: 47912
diff changeset
   262
                             false true [] @{prop False}
45551
a62c7a21f4ab removed needless baggage
blanchet
parents: 45508
diff changeset
   263
      |> #1
43566
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
   264
    val atp_problem =
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
   265
      atp_problem
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   266
      |> map (apsnd (filter_out (is_problem_line_tautology ctxt format)))
43576
ebeda6275027 don't export any metastrange or other nonatomizable formulas, since these don't help proving normal things, they are somewhat broken in the ATP output, and they are atypical
blanchet
parents: 43572
diff changeset
   267
    val all_names = facts |> map (Thm.get_name_hint o snd)
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   268
    val infers =
43576
ebeda6275027 don't export any metastrange or other nonatomizable formulas, since these don't help proving normal things, they are somewhat broken in the ATP output, and they are atypical
blanchet
parents: 43572
diff changeset
   269
      facts |> map (fn (_, th) =>
ebeda6275027 don't export any metastrange or other nonatomizable formulas, since these don't help proving normal things, they are somewhat broken in the ATP output, and they are atypical
blanchet
parents: 43572
diff changeset
   270
                       (fact_name_of (Thm.get_name_hint th),
ebeda6275027 don't export any metastrange or other nonatomizable formulas, since these don't help proving normal things, they are somewhat broken in the ATP output, and they are atypical
blanchet
parents: 43572
diff changeset
   271
                        theorems_mentioned_in_proof_term (SOME all_names) th))
43479
5af1abc13c1f only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
blanchet
parents: 43468
diff changeset
   272
    val all_atp_problem_names =
5af1abc13c1f only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
blanchet
parents: 43468
diff changeset
   273
      atp_problem |> maps (map ident_of_problem_line o snd)
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   274
    val infers =
43499
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   275
      infers |> filter (member (op =) all_atp_problem_names o fst)
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   276
             |> map (apsnd (filter (member (op =) all_atp_problem_names)))
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   277
    val ordered_names =
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   278
      String_Graph.empty
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   279
      |> fold (String_Graph.new_node o rpair ()) all_atp_problem_names
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   280
      |> fold (fn (to, froms) =>
43566
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
   281
                  fold (fn from => String_Graph.add_edge (from, to)) froms)
a818d5a34cca filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents: 43499
diff changeset
   282
              infers
43499
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   283
      |> String_Graph.topological_order
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   284
    val order_tab =
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   285
      Symtab.empty
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   286
      |> fold (Symtab.insert (op =))
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   287
              (ordered_names ~~ (1 upto length ordered_names))
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   288
    val name_ord = int_ord o pairself (the o Symtab.lookup order_tab)
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   289
    val atp_problem =
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   290
      atp_problem
48131
1016664b8feb started adding polymophic SPASS output
blanchet
parents: 48130
diff changeset
   291
      |> (case format of DFG _ => I | _ => add_inferences_to_problem infers)
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   292
      |> order_problem_facts name_ord
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 46734
diff changeset
   293
    val ord = effective_term_order ctxt eN (* dummy *)
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 46734
diff changeset
   294
    val ss = lines_for_atp_problem format ord (K []) atp_problem
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   295
    val _ = app (File.append path) ss
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   296
  in () end
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   297
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   298
end;