src/HOL/TPTP/atp_theory_export.ML
author blanchet
Wed, 18 Jul 2012 08:44:04 +0200
changeset 48318 325c8fd0d762
parent 48316 252f45c04042
child 48324 3ee5b5589402
permissions -rw-r--r--
more consolidation of MaSh code
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
48234
06216c789ac9 moved MaSh into own files
blanchet
parents: 48233
diff changeset
     5
Export Isabelle theories as first-order TPTP inferences.
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
     6
*)
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
     7
46321
484dc68c8c89 renamed theory exporter
blanchet
parents: 46320
diff changeset
     8
signature ATP_THEORY_EXPORT =
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
     9
sig
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
    10
  type atp_format = ATP_Problem.atp_format
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
    11
48234
06216c789ac9 moved MaSh into own files
blanchet
parents: 48233
diff changeset
    12
  val generate_atp_inference_file_for_theory :
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
    13
    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
    14
end;
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    15
48234
06216c789ac9 moved MaSh into own files
blanchet
parents: 48233
diff changeset
    16
structure ATP_Theory_Export : ATP_THEORY_EXPORT =
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    17
struct
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    18
43479
5af1abc13c1f only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
blanchet
parents: 43468
diff changeset
    19
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
    20
open ATP_Proof
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 46301
diff changeset
    21
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
    22
open ATP_Systems
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    23
48234
06216c789ac9 moved MaSh into own files
blanchet
parents: 48233
diff changeset
    24
val fact_name_of = prefix fact_prefix o ascii_of
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    25
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    26
fun inference_term [] = NONE
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    27
  | inference_term ss =
48132
9aa0fad4e864 added type arguments to "ATerm" constructor -- but don't use them yet
blanchet
parents: 48131
diff changeset
    28
    ATerm (("inference", []),
9aa0fad4e864 added type arguments to "ATerm" constructor -- but don't use them yet
blanchet
parents: 48131
diff changeset
    29
           [ATerm (("isabelle", []), []),
9aa0fad4e864 added type arguments to "ATerm" constructor -- but don't use them yet
blanchet
parents: 48131
diff changeset
    30
            ATerm ((tptp_empty_list, []), []),
9aa0fad4e864 added type arguments to "ATerm" constructor -- but don't use them yet
blanchet
parents: 48131
diff changeset
    31
            ATerm ((tptp_empty_list, []),
9aa0fad4e864 added type arguments to "ATerm" constructor -- but don't use them yet
blanchet
parents: 48131
diff changeset
    32
            map (fn s => ATerm ((s, []), [])) ss)])
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    33
    |> SOME
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    34
fun inference infers ident =
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    35
  these (AList.lookup (op =) infers ident) |> inference_term
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    36
fun add_inferences_to_problem_line infers
46406
0e490b9e8422 extended SPASS/DFG output with ranks
blanchet
parents: 46365
diff changeset
    37
                                   (Formula (ident, Axiom, phi, NONE, tms)) =
0e490b9e8422 extended SPASS/DFG output with ranks
blanchet
parents: 46365
diff changeset
    38
    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
    39
  | add_inferences_to_problem_line _ line = line
43996
4d1270ddf042 make SML/NJ happy
blanchet
parents: 43863
diff changeset
    40
fun add_inferences_to_problem infers =
4d1270ddf042 make SML/NJ happy
blanchet
parents: 43863
diff changeset
    41
  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
    42
48142
efaff8206967 finished implementation of DFG type class output
blanchet
parents: 48140
diff changeset
    43
fun ident_of_problem_line (Class_Decl (ident, _, _)) = ident
efaff8206967 finished implementation of DFG type class output
blanchet
parents: 48140
diff changeset
    44
  | ident_of_problem_line (Type_Decl (ident, _, _)) = ident
48137
6f524f2066e3 cleanly distinguish between type declarations and symbol declarations
blanchet
parents: 48132
diff changeset
    45
  | ident_of_problem_line (Sym_Decl (ident, _, _)) = ident
48142
efaff8206967 finished implementation of DFG type class output
blanchet
parents: 48140
diff changeset
    46
  | 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
    47
  | 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
    48
48318
325c8fd0d762 more consolidation of MaSh code
blanchet
parents: 48316
diff changeset
    49
fun atp_for_format (THF (Polymorphic, _, _, _)) = dummy_thfN
325c8fd0d762 more consolidation of MaSh code
blanchet
parents: 48316
diff changeset
    50
  | atp_for_format (THF (Monomorphic, _, _, _)) = satallaxN
325c8fd0d762 more consolidation of MaSh code
blanchet
parents: 48316
diff changeset
    51
  | atp_for_format (DFG Polymorphic) = spass_polyN
325c8fd0d762 more consolidation of MaSh code
blanchet
parents: 48316
diff changeset
    52
  | atp_for_format (DFG Monomorphic) = spassN
325c8fd0d762 more consolidation of MaSh code
blanchet
parents: 48316
diff changeset
    53
  | atp_for_format (TFF (Polymorphic, _)) = alt_ergoN
325c8fd0d762 more consolidation of MaSh code
blanchet
parents: 48316
diff changeset
    54
  | atp_for_format (TFF (Monomorphic, _)) = vampireN
325c8fd0d762 more consolidation of MaSh code
blanchet
parents: 48316
diff changeset
    55
  | atp_for_format FOF = eN
325c8fd0d762 more consolidation of MaSh code
blanchet
parents: 48316
diff changeset
    56
  | atp_for_format CNF_UEQ = waldmeisterN
325c8fd0d762 more consolidation of MaSh code
blanchet
parents: 48316
diff changeset
    57
  | atp_for_format CNF = eN
325c8fd0d762 more consolidation of MaSh code
blanchet
parents: 48316
diff changeset
    58
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
    59
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
    60
  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
    61
    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
    62
    val prob_file = File.tmp_path (Path.explode "prob")
48318
325c8fd0d762 more consolidation of MaSh code
blanchet
parents: 48316
diff changeset
    63
    val atp = atp_for_format format
47606
06dde48a1503 true delayed evaluation of "SPASS_VERSION" environment variable
blanchet
parents: 47055
diff changeset
    64
    val {exec, arguments, proof_delims, known_failures, ...} =
06dde48a1503 true delayed evaluation of "SPASS_VERSION" environment variable
blanchet
parents: 47055
diff changeset
    65
      get_atp thy atp ()
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 46734
diff changeset
    66
    val ord = effective_term_order ctxt atp
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 46734
diff changeset
    67
    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
    68
                    |> File.write_list prob_file
48213
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    69
    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
    70
    val command =
48213
d20add034f64 first go at generating files for MaSh (machine-learning Sledgehammer)
blanchet
parents: 48142
diff changeset
    71
      File.shell_path (Path.explode path) ^
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 46734
diff changeset
    72
      " " ^ 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
    73
      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
    74
  in
43850
7f2cbc713344 moved bash operations to Isabelle_System (cf. Scala version);
wenzelm
parents: 43828
diff changeset
    75
    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
    76
    |> fst
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 46734
diff changeset
    77
    |> 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
    78
    |> 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
    79
  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
    80
  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
    81
48217
8994afe09c18 more precise dependencies -- eliminate tautologies
blanchet
parents: 48216
diff changeset
    82
val 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
    83
  [@{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
    84
  |> 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
    85
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
    86
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
    87
    exists (fn prefix => String.isPrefix prefix ident)
48217
8994afe09c18 more precise dependencies -- eliminate tautologies
blanchet
parents: 48216
diff changeset
    88
           tautology_prefixes andalso
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
    89
    is_none (run_some_atp ctxt format
46406
0e490b9e8422 extended SPASS/DFG output with ranks
blanchet
parents: 46365
diff changeset
    90
                 [(factsN, [Formula (ident, Conjecture, phi, NONE, [])])])
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
    91
  | 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
    92
43499
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
    93
fun order_facts ord = sort (ord o pairself ident_of_problem_line)
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
    94
fun order_problem_facts _ [] = []
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
    95
  | order_problem_facts ord ((heading, lines) :: problem) =
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
    96
    if heading = factsN then (heading, order_facts ord lines) :: problem
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
    97
    else (heading, lines) :: order_problem_facts ord problem
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
    98
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
    99
(* A fairly random selection of types used for monomorphizing. *)
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   100
val ground_types =
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   101
  [@{typ nat}, HOLogic.intT, HOLogic.realT, @{typ "nat => bool"}, @{typ bool},
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   102
   @{typ unit}]
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   103
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   104
fun ground_type_for_tvar _ [] tvar =
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   105
    raise TYPE ("ground_type_for_sorts", [TVar tvar], [])
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   106
  | ground_type_for_tvar thy (T :: Ts) tvar =
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   107
    if can (Sign.typ_match thy (TVar tvar, T)) Vartab.empty then T
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   108
    else ground_type_for_tvar thy Ts tvar
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   109
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   110
fun monomorphize_term ctxt t =
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   111
  let val thy = Proof_Context.theory_of ctxt in
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   112
    t |> map_types (map_type_tvar (ground_type_for_tvar thy ground_types))
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   113
    handle TYPE _ => @{prop True}
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   114
  end
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   115
48318
325c8fd0d762 more consolidation of MaSh code
blanchet
parents: 48316
diff changeset
   116
fun all_non_tautological_facts_of ctxt prover thy css_table =
48315
82d6e46c673f fixed order of accessibles + other tweaks to MaSh
blanchet
parents: 48301
diff changeset
   117
  Sledgehammer_Fact.all_facts_of thy css_table
48318
325c8fd0d762 more consolidation of MaSh code
blanchet
parents: 48316
diff changeset
   118
  |> filter_out ((Sledgehammer_Filter_MaSh.is_likely_tautology ctxt prover orf
48315
82d6e46c673f fixed order of accessibles + other tweaks to MaSh
blanchet
parents: 48301
diff changeset
   119
                  Sledgehammer_Filter_MaSh.is_too_meta) o snd)
82d6e46c673f fixed order of accessibles + other tweaks to MaSh
blanchet
parents: 48301
diff changeset
   120
48234
06216c789ac9 moved MaSh into own files
blanchet
parents: 48233
diff changeset
   121
fun generate_atp_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
   122
  let
48318
325c8fd0d762 more consolidation of MaSh code
blanchet
parents: 48316
diff changeset
   123
    val atp = atp_for_format format
48299
5e5c6616f0fe centrally construct expensive data structures
blanchet
parents: 48251
diff changeset
   124
    val css_table = Sledgehammer_Fact.clasimpset_rule_table_of ctxt
48315
82d6e46c673f fixed order of accessibles + other tweaks to MaSh
blanchet
parents: 48301
diff changeset
   125
    val type_enc =
82d6e46c673f fixed order of accessibles + other tweaks to MaSh
blanchet
parents: 48301
diff changeset
   126
      type_enc |> type_enc_from_string Strict
82d6e46c673f fixed order of accessibles + other tweaks to MaSh
blanchet
parents: 48301
diff changeset
   127
               |> adjust_type_enc format
48131
1016664b8feb started adding polymophic SPASS output
blanchet
parents: 48130
diff changeset
   128
    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
   129
    val path = file_name |> Path.explode
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   130
    val _ = File.write path ""
48318
325c8fd0d762 more consolidation of MaSh code
blanchet
parents: 48316
diff changeset
   131
    val facts = all_non_tautological_facts_of ctxt atp thy css_table
45551
a62c7a21f4ab removed needless baggage
blanchet
parents: 45508
diff changeset
   132
    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
   133
      facts
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   134
      |> map (fn ((_, loc), th) =>
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   135
                 ((Thm.get_name_hint th, loc),
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   136
                   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
   137
      |> 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
   138
                             false true [] @{prop False}
45551
a62c7a21f4ab removed needless baggage
blanchet
parents: 45508
diff changeset
   139
      |> #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
   140
    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
   141
      atp_problem
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   142
      |> map (apsnd (filter_out (is_problem_line_tautology ctxt format)))
48217
8994afe09c18 more precise dependencies -- eliminate tautologies
blanchet
parents: 48216
diff changeset
   143
    val ths = facts |> map snd
48316
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
   144
    val all_names = ths |> map (rpair () o Thm.get_name_hint) |> Symtab.make
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   145
    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
   146
      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
   147
                       (fact_name_of (Thm.get_name_hint th),
48251
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   148
                        th |> Sledgehammer_Util.thms_in_proof (SOME all_names)
48234
06216c789ac9 moved MaSh into own files
blanchet
parents: 48233
diff changeset
   149
                           |> map fact_name_of))
43479
5af1abc13c1f only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
blanchet
parents: 43468
diff changeset
   150
    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
   151
      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
   152
    val infers =
43499
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   153
      infers |> filter (member (op =) all_atp_problem_names o fst)
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   154
             |> map (apsnd (filter (member (op =) all_atp_problem_names)))
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   155
    val ordered_names =
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   156
      String_Graph.empty
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   157
      |> fold (String_Graph.new_node o rpair ()) all_atp_problem_names
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   158
      |> 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
   159
                  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
   160
              infers
43499
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   161
      |> String_Graph.topological_order
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   162
    val order_tab =
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   163
      Symtab.empty
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   164
      |> fold (Symtab.insert (op =))
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   165
              (ordered_names ~~ (1 upto length ordered_names))
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   166
    val name_ord = int_ord o pairself (the o Symtab.lookup order_tab)
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   167
    val atp_problem =
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   168
      atp_problem
48131
1016664b8feb started adding polymophic SPASS output
blanchet
parents: 48130
diff changeset
   169
      |> (case format of DFG _ => I | _ => add_inferences_to_problem infers)
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   170
      |> order_problem_facts name_ord
47038
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 46734
diff changeset
   171
    val ord = effective_term_order ctxt eN (* dummy *)
2409b484e1cc continued implementation of term ordering attributes
blanchet
parents: 46734
diff changeset
   172
    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
   173
    val _ = app (File.append path) ss
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   174
  in () end
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   175
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   176
end;