src/HOL/TPTP/atp_theory_export.ML
author wenzelm
Mon, 11 Sep 2023 19:30:48 +0200
changeset 78659 b5f3d1051b13
parent 77429 110988ad5b4c
child 80306 c2537860ccf8
permissions -rw-r--r--
tuned;
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
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    12
  datatype inference_policy =
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    13
    No_Inferences | Unchecked_Inferences | Checked_Inferences
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    14
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
    15
  val generate_atp_inference_file_for_theory : Proof.context -> theory -> atp_format ->
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
    16
    inference_policy -> string -> string -> unit
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
    17
  val generate_casc_lbt_isa_files_for_theory : Proof.context -> theory -> atp_format ->
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
    18
    inference_policy -> string -> string -> unit
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    19
end;
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    20
48234
06216c789ac9 moved MaSh into own files
blanchet
parents: 48233
diff changeset
    21
structure ATP_Theory_Export : ATP_THEORY_EXPORT =
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    22
struct
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    23
43479
5af1abc13c1f only refer to facts found in TPTP file -- e.g. facts that simplify to true are excluded
blanchet
parents: 43468
diff changeset
    24
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
    25
open ATP_Proof
46320
0b8b73b49848 renamed two files to make room for a new file
blanchet
parents: 46301
diff changeset
    26
open ATP_Problem_Generate
72400
abfeed05c323 tune filename
desharna
parents: 70943
diff changeset
    27
open Sledgehammer_ATP_Systems
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    28
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
    29
val max_dependencies = 100
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
    30
val max_facts = 512
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
    31
val atp_timeout = seconds 0.5
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
    32
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    33
datatype inference_policy =
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    34
  No_Inferences | Unchecked_Inferences | Checked_Inferences
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    35
53980
7e6a82c593f4 made SML/NJ happy (NB: toplevel ML environment is unmanaged);
wenzelm
parents: 53586
diff changeset
    36
val prefix = Library.prefix
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    37
val fact_name_of = prefix fact_prefix o ascii_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
    38
74902
ece4f07ebb04 fixed HOL-TPTP
desharna
parents: 74044
diff changeset
    39
fun atp_of_format (THF (Polymorphic, _, _)) = leo3N
ece4f07ebb04 fixed HOL-TPTP
desharna
parents: 74044
diff changeset
    40
  | atp_of_format (THF (Monomorphic, _, _)) = satallaxN
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51652
diff changeset
    41
  | atp_of_format (DFG Monomorphic) = spassN
74902
ece4f07ebb04 fixed HOL-TPTP
desharna
parents: 74044
diff changeset
    42
  | atp_of_format (TFF (Polymorphic, _)) = alt_ergoN
ece4f07ebb04 fixed HOL-TPTP
desharna
parents: 74044
diff changeset
    43
  | atp_of_format (TFF (Monomorphic, _)) = vampireN
70943
blanchet
parents: 69597
diff changeset
    44
  | atp_of_format FOF = eN (* FIXME? *)
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51652
diff changeset
    45
  | atp_of_format CNF_UEQ = waldmeisterN
70943
blanchet
parents: 69597
diff changeset
    46
  | atp_of_format CNF = eN (* FIXME? *)
48318
325c8fd0d762 more consolidation of MaSh code
blanchet
parents: 48316
diff changeset
    47
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    48
fun run_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
    49
  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
    50
    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
    51
    val prob_file = File.tmp_path (Path.explode "prob")
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51652
diff changeset
    52
    val atp = atp_of_format format
57676
d53b1f876afb compile
blanchet
parents: 57307
diff changeset
    53
    val {exec, arguments, proof_delims, known_failures, ...} = get_atp thy atp ()
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
    54
    val _ = problem
75344
647611e6da76 compile TPTP module
blanchet
parents: 74902
diff changeset
    55
      |> lines_of_atp_problem format (K [])
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
    56
      |> File.write_list prob_file
48376
416e4123baf3 use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents: 48324
diff changeset
    57
    val path = getenv (List.last (fst exec)) ^ "/" ^ List.last (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
    58
    val command =
73575
23d2adc5489e compile;
wenzelm
parents: 73433
diff changeset
    59
      space_implode " " (File.bash_path (Path.explode path) ::
77429
110988ad5b4c compile
blanchet
parents: 75344
diff changeset
    60
        arguments false false "" atp_timeout prob_file)
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    61
    val outcome =
62519
a564458f94db tuned signature -- clarified modules;
wenzelm
parents: 61860
diff changeset
    62
      Timeout.apply atp_timeout Isabelle_System.bash_output command
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    63
      |> fst
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    64
      |> extract_tstplike_proof_and_outcome false proof_delims known_failures
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    65
      |> snd
62519
a564458f94db tuned signature -- clarified modules;
wenzelm
parents: 61860
diff changeset
    66
      handle Timeout.TIMEOUT _ => SOME TimedOut
51647
25acf689a53e no need to filter tautologies anymore -- they are prefiltered by "all_facts"'
blanchet
parents: 51646
diff changeset
    67
    val _ =
25acf689a53e no need to filter tautologies anymore -- they are prefiltered by "all_facts"'
blanchet
parents: 51646
diff changeset
    68
      tracing ("Ran ATP: " ^
51649
5d882158c221 compile + fixed naming convention
blanchet
parents: 51647
diff changeset
    69
               (case outcome of
5d882158c221 compile + fixed naming convention
blanchet
parents: 51647
diff changeset
    70
                  NONE => "Success"
53586
bd5fa6425993 prefixed types and some functions with "atp_" for disambiguation
blanchet
parents: 52995
diff changeset
    71
                | SOME failure => string_of_atp_failure failure))
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    72
  in outcome end
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
61860
2ce3d12015b3 cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
blanchet
parents: 61323
diff changeset
    74
fun is_problem_line_reprovable ctxt format prelude axioms deps (Formula (ident', _, phi, _, _)) =
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    75
    is_none (run_atp ctxt format
61860
2ce3d12015b3 cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
blanchet
parents: 61323
diff changeset
    76
      ((factsN,
2ce3d12015b3 cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
blanchet
parents: 61323
diff changeset
    77
        Formula (ident', Conjecture, phi, NONE, []) :: map_filter (Symtab.lookup axioms) deps) ::
2ce3d12015b3 cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
blanchet
parents: 61323
diff changeset
    78
       prelude))
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    79
  | is_problem_line_reprovable _ _ _ _ _ _ = false
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    80
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    81
fun inference_term _ [] = NONE
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    82
  | inference_term check_infs ss =
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    83
    ATerm (("inference", []),
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    84
        [ATerm (("checked_isabelle" |> not check_infs ? prefix "un", []), []),
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    85
         ATerm ((tptp_empty_list, []), []),
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    86
         ATerm ((tptp_empty_list, []),
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    87
         map (fn s => ATerm ((s, []), [])) ss)])
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    88
    |> SOME
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    89
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    90
fun add_inferences_to_problem_line ctxt format check_infs prelude axioms infers
61860
2ce3d12015b3 cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
blanchet
parents: 61323
diff changeset
    91
      (line as Formula ((ident, alt), Axiom, phi, NONE, info)) =
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    92
    let
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    93
      val deps =
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    94
        case these (AList.lookup (op =) infers ident) of
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    95
          [] => []
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    96
        | deps =>
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    97
          if check_infs andalso
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    98
             not (is_problem_line_reprovable ctxt format prelude axioms deps
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
    99
                                             line) then
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   100
            []
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   101
          else
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   102
            deps
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   103
    in
61860
2ce3d12015b3 cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
blanchet
parents: 61323
diff changeset
   104
      Formula ((ident, alt), Lemma, phi, inference_term check_infs deps, info)
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   105
    end
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   106
  | add_inferences_to_problem_line _ _ _ _ _ _ line = line
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   107
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   108
fun add_inferences_to_problem ctxt format check_infs prelude infers problem =
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   109
  let
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   110
    fun add_if_axiom (axiom as Formula ((ident, _), Axiom, _, _, _)) = Symtab.default (ident, axiom)
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   111
      | add_if_axiom _ = I
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   112
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   113
    val add_axioms_of_problem = fold (fold add_if_axiom o snd)
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   114
    val axioms = Symtab.empty |> check_infs ? add_axioms_of_problem problem
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   115
  in
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   116
    map (apsnd (map (add_inferences_to_problem_line ctxt format check_infs prelude axioms infers)))
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   117
      problem
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   118
  end
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   119
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   120
fun ident_of_problem_line (Class_Decl (ident, _, _)) = ident
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   121
  | ident_of_problem_line (Type_Decl (ident, _, _)) = ident
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   122
  | ident_of_problem_line (Sym_Decl (ident, _, _)) = ident
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   123
  | ident_of_problem_line (Class_Memb (ident, _, _, _)) = ident
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   124
  | ident_of_problem_line (Formula ((ident, _), _, _, _, _)) = ident
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   125
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58922
diff changeset
   126
fun order_facts ord = sort (ord o apply2 ident_of_problem_line)
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   127
43499
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   128
fun order_problem_facts _ [] = []
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   129
  | order_problem_facts ord ((heading, lines) :: problem) =
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   130
    if heading = factsN then (heading, order_facts ord lines) :: problem
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   131
    else (heading, lines) :: order_problem_facts ord problem
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   132
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   133
(* A fairly random selection of types used for monomorphizing. *)
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   134
val ground_types =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62549
diff changeset
   135
  [\<^typ>\<open>nat\<close>, HOLogic.intT, HOLogic.realT, \<^typ>\<open>nat => bool\<close>, \<^typ>\<open>bool\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62549
diff changeset
   136
   \<^typ>\<open>unit\<close>]
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   137
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   138
fun ground_type_of_tvar _ [] tvar = raise TYPE ("ground_type_of_tvar", [TVar tvar], [])
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51652
diff changeset
   139
  | ground_type_of_tvar thy (T :: Ts) tvar =
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   140
    if can (Sign.typ_match thy (TVar tvar, T)) Vartab.empty then T
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51652
diff changeset
   141
    else ground_type_of_tvar thy Ts tvar
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   142
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   143
fun monomorphize_term ctxt t =
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   144
  let val thy = Proof_Context.theory_of ctxt in
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51652
diff changeset
   145
    t |> map_types (map_type_tvar (ground_type_of_tvar thy ground_types))
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62549
diff changeset
   146
    handle TYPE _ => \<^prop>\<open>True\<close>
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   147
  end
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   148
51633
f11a1498dfdc try to preserve original linearization
blanchet
parents: 51632
diff changeset
   149
fun heading_sort_key heading =
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   150
  if String.isPrefix factsN heading then "_" ^ heading else heading
57306
ff10067b2248 optimized traversal of proof terms by skipping bad apples (e.g. full_exhaustive_int'.pinduct)
blanchet
parents: 57268
diff changeset
   151
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51652
diff changeset
   152
fun problem_of_theory ctxt thy format infer_policy type_enc =
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   153
  let
48299
5e5c6616f0fe centrally construct expensive data structures
blanchet
parents: 48251
diff changeset
   154
    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
   155
    val type_enc =
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51998
diff changeset
   156
      type_enc |> type_enc_of_string Non_Strict
48315
82d6e46c673f fixed order of accessibles + other tweaks to MaSh
blanchet
parents: 48301
diff changeset
   157
               |> adjust_type_enc format
48131
1016664b8feb started adding polymophic SPASS output
blanchet
parents: 48130
diff changeset
   158
    val mono = not (is_type_enc_polymorphic type_enc)
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   159
48438
3e45c98fe127 distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
blanchet
parents: 48406
diff changeset
   160
    val facts =
73942
57423714c29d fixed HOL-TPTP following f58108b7a60c
desharna
parents: 73575
diff changeset
   161
      Sledgehammer_Fact.all_facts (Proof_Context.init_global thy) true Keyword.empty_keywords [] []
57423714c29d fixed HOL-TPTP following f58108b7a60c
desharna
parents: 73575
diff changeset
   162
        css_table
60641
f6e8293747fd clarified context;
wenzelm
parents: 59582
diff changeset
   163
      |> sort (Sledgehammer_MaSh.crude_thm_ord ctxt o apply2 snd)
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   164
    val 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
   165
      facts
45305
3e09961326ce also export DFG formats
blanchet
parents: 45301
diff changeset
   166
      |> map (fn ((_, loc), th) =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59577
diff changeset
   167
        ((Thm.get_name_hint th, loc), th |> Thm.prop_of |> mono ? monomorphize_term ctxt))
61860
2ce3d12015b3 cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
blanchet
parents: 61323
diff changeset
   168
      |> generate_atp_problem ctxt true format Axiom type_enc Exporter combsN false false true []
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62549
diff changeset
   169
        \<^prop>\<open>False\<close>
60924
610794dff23c tuned signature, in accordance to sortBy in Scala;
wenzelm
parents: 60641
diff changeset
   170
      |> #1 |> sort_by (heading_sort_key o fst)
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   171
    val prelude = fst (split_last problem)
50735
6b232d76cbc9 refined class handling, to prevent cycles in fact graph
blanchet
parents: 50521
diff changeset
   172
    val name_tabs = Sledgehammer_Fact.build_name_tables Thm.get_name_hint facts
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   173
    val infers =
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   174
      if infer_policy = No_Inferences then
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   175
        []
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   176
      else
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   177
        facts
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   178
        |> map (fn (_, th) =>
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   179
                   (fact_name_of (Thm.get_name_hint th),
57306
ff10067b2248 optimized traversal of proof terms by skipping bad apples (e.g. full_exhaustive_int'.pinduct)
blanchet
parents: 57268
diff changeset
   180
                    th |> Sledgehammer_Util.thms_in_proof max_dependencies (SOME name_tabs)
ff10067b2248 optimized traversal of proof terms by skipping bad apples (e.g. full_exhaustive_int'.pinduct)
blanchet
parents: 57268
diff changeset
   181
                       |> these |> map fact_name_of))
51650
3dd495cd98a2 smoothly handle cyclic graphs
blanchet
parents: 51649
diff changeset
   182
    val all_problem_names =
57307
blanchet
parents: 57306
diff changeset
   183
      problem |> maps (map ident_of_problem_line o snd) |> distinct (op =)
51650
3dd495cd98a2 smoothly handle cyclic graphs
blanchet
parents: 51649
diff changeset
   184
    val all_problem_name_set = Symtab.make (map (rpair ()) all_problem_names)
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   185
    val infers =
51650
3dd495cd98a2 smoothly handle cyclic graphs
blanchet
parents: 51649
diff changeset
   186
      infers |> filter (Symtab.defined all_problem_name_set o fst)
3dd495cd98a2 smoothly handle cyclic graphs
blanchet
parents: 51649
diff changeset
   187
             |> map (apsnd (filter (Symtab.defined all_problem_name_set)))
3dd495cd98a2 smoothly handle cyclic graphs
blanchet
parents: 51649
diff changeset
   188
    val maybe_add_edge = perhaps o try o String_Graph.add_edge_acyclic
43499
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   189
    val ordered_names =
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   190
      String_Graph.empty
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   191
      |> fold (String_Graph.new_node o rpair ()) all_problem_names
57307
blanchet
parents: 57306
diff changeset
   192
      |> fold (fn (to, froms) => fold (fn from => maybe_add_edge (from, to)) froms) infers
51650
3dd495cd98a2 smoothly handle cyclic graphs
blanchet
parents: 51649
diff changeset
   193
      |> fold (fn (to, from) => maybe_add_edge (from, to))
57307
blanchet
parents: 57306
diff changeset
   194
        (tl all_problem_names ~~ fst (split_last all_problem_names))
43499
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   195
      |> String_Graph.topological_order
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   196
    val order_tab =
9ca694caa61b order generated facts topologically
blanchet
parents: 43498
diff changeset
   197
      Symtab.empty
57307
blanchet
parents: 57306
diff changeset
   198
      |> fold (Symtab.insert (op =)) (ordered_names ~~ (1 upto length ordered_names))
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58922
diff changeset
   199
    val name_ord = int_ord o apply2 (the o Symtab.lookup order_tab)
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   200
  in
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   201
    (facts,
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   202
     problem
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   203
     |> (case format of
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   204
          DFG _ => I
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   205
        | _ => add_inferences_to_problem ctxt format (infer_policy = Checked_Inferences) prelude
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   206
          infers)
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   207
     |> order_problem_facts name_ord)
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   208
  end
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   209
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   210
fun write_lines path ss =
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   211
  let
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   212
    val _ = File.write path ""
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   213
    val _ = app (File.append path) ss
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   214
  in () end
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   215
70943
blanchet
parents: 69597
diff changeset
   216
fun generate_atp_inference_file_for_theory ctxt thy format infer_policy type_enc file_name =
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   217
  let
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   218
    val (_, problem) = problem_of_theory ctxt thy format infer_policy type_enc
75344
647611e6da76 compile TPTP module
blanchet
parents: 74902
diff changeset
   219
    val ss = lines_of_atp_problem format (K []) problem
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   220
  in write_lines (Path.explode file_name) ss end
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   221
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   222
fun ap dir = Path.append dir o Path.explode
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   223
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   224
fun chop_maximal_groups eq xs =
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   225
  let
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   226
    val rev_xs = rev xs
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   227
    fun chop_group _ [] = []
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   228
      | chop_group n (xs as x :: _) =
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   229
        let
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   230
          val n' = find_index (curry eq x) rev_xs
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   231
          val (ws', xs') = chop (n - n') xs
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   232
        in ws' :: chop_group n' xs' end
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   233
   in chop_group (length xs) xs end
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   234
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   235
fun theory_name_of_fact (Formula ((_, alt), _, _, _, _)) =
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   236
    (case first_field Long_Name.separator alt of
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   237
       NONE => alt
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   238
     | SOME (thy, _) => thy)
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   239
  | theory_name_of_fact _ = ""
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   240
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   241
val problem_suffix = ".p"
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   242
val suggestion_suffix = ".sugg"
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   243
val include_suffix = ".ax"
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   244
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   245
val file_order_name = "FilesInProcessingOrder"
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   246
val problem_order_name = "ProblemsInProcessingOrder"
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   247
val problem_name = "problems"
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   248
val suggestion_name = "suggestions"
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   249
val include_name = "incl"
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   250
val prelude_base_name = "prelude"
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   251
val prelude_name = prelude_base_name ^ include_suffix
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   252
51652
5ff01d585a8c handle case clashes on Mac file system by encoding goal numbers
blanchet
parents: 51650
diff changeset
   253
val encode_meta = Sledgehammer_MaSh.encode_str
5ff01d585a8c handle case clashes on Mac file system by encoding goal numbers
blanchet
parents: 51650
diff changeset
   254
53980
7e6a82c593f4 made SML/NJ happy (NB: toplevel ML environment is unmanaged);
wenzelm
parents: 53586
diff changeset
   255
fun include_base_name_of_fact x = encode_meta (theory_name_of_fact x)
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   256
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   257
fun include_line base_name =
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   258
  "include('" ^ include_name ^ "/" ^ base_name ^ include_suffix ^ "').\n"
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   259
51652
5ff01d585a8c handle case clashes on Mac file system by encoding goal numbers
blanchet
parents: 51650
diff changeset
   260
val hol_base_name = encode_meta "HOL"
5ff01d585a8c handle case clashes on Mac file system by encoding goal numbers
blanchet
parents: 51650
diff changeset
   261
5ff01d585a8c handle case clashes on Mac file system by encoding goal numbers
blanchet
parents: 51650
diff changeset
   262
fun should_generate_problem thy base_name (Formula ((_, alt), _, _, _, _)) =
61860
2ce3d12015b3 cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
blanchet
parents: 61323
diff changeset
   263
  (case try (Global_Theory.get_thm thy) alt of
51652
5ff01d585a8c handle case clashes on Mac file system by encoding goal numbers
blanchet
parents: 51650
diff changeset
   264
    SOME th =>
61860
2ce3d12015b3 cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
blanchet
parents: 61323
diff changeset
   265
    (* This is a crude hack to detect theorems stated and proved by the user (as opposed to those
2ce3d12015b3 cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
blanchet
parents: 61323
diff changeset
   266
       derived by various packages). In addition, we leave out everything in "HOL" as too basic to
2ce3d12015b3 cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
blanchet
parents: 61323
diff changeset
   267
       be interesting. *)
51652
5ff01d585a8c handle case clashes on Mac file system by encoding goal numbers
blanchet
parents: 51650
diff changeset
   268
    Thm.legacy_get_kind th <> "" andalso base_name <> hol_base_name
61860
2ce3d12015b3 cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
blanchet
parents: 61323
diff changeset
   269
  | NONE => false)
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   270
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   271
(* Convention: theoryname__goalname *)
51652
5ff01d585a8c handle case clashes on Mac file system by encoding goal numbers
blanchet
parents: 51650
diff changeset
   272
fun problem_name_of base_name n alt =
5ff01d585a8c handle case clashes on Mac file system by encoding goal numbers
blanchet
parents: 51650
diff changeset
   273
  base_name ^ "__" ^ string_of_int n ^ "_" ^
5ff01d585a8c handle case clashes on Mac file system by encoding goal numbers
blanchet
parents: 51650
diff changeset
   274
  perhaps (try (unprefix (base_name ^ "_"))) alt ^ problem_suffix
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   275
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   276
fun suggestion_name_of base_name n alt =
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   277
  base_name ^ "__" ^ string_of_int n ^ "_" ^
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   278
  perhaps (try (unprefix (base_name ^ "_"))) alt ^ suggestion_suffix
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   279
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   280
fun generate_casc_lbt_isa_files_for_theory ctxt thy format infer_policy type_enc dir_name =
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   281
  let
73315
wenzelm
parents: 72591
diff changeset
   282
    val dir = Isabelle_System.make_directory (Path.explode dir_name)
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   283
    val file_order_path = ap dir file_order_name
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   284
    val _ = File.write file_order_path ""
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   285
    val problem_order_path = ap dir problem_order_name
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   286
    val _ = File.write problem_order_path ""
73315
wenzelm
parents: 72591
diff changeset
   287
    val problem_dir = Isabelle_System.make_directory (ap dir problem_name)
wenzelm
parents: 72591
diff changeset
   288
    val suggestion_dir = Isabelle_System.make_directory (ap dir suggestion_name)
wenzelm
parents: 72591
diff changeset
   289
    val include_dir = Isabelle_System.make_directory (ap problem_dir include_name)
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   290
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   291
    val (facts, (prelude, groups)) =
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51652
diff changeset
   292
      problem_of_theory ctxt thy format infer_policy type_enc
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   293
      ||> split_last
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   294
      ||> apsnd (snd
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58922
diff changeset
   295
           #> chop_maximal_groups (op = o apply2 theory_name_of_fact)
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   296
           #> map (`(include_base_name_of_fact o hd)))
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   297
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   298
    val fact_tab = Symtab.make (map (fn fact as (_, th) => (Thm.get_name_hint th, fact)) facts)
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   299
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   300
    fun write_prelude () =
75344
647611e6da76 compile TPTP module
blanchet
parents: 74902
diff changeset
   301
      let val ss = lines_of_atp_problem format (K []) prelude in
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   302
        File.append file_order_path (prelude_base_name ^ "\n");
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   303
        write_lines (ap include_dir prelude_name) ss
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   304
      end
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   305
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   306
    fun write_include_file (base_name, fact_lines) =
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   307
      let
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   308
        val name = base_name ^ include_suffix
75344
647611e6da76 compile TPTP module
blanchet
parents: 74902
diff changeset
   309
        val ss = lines_of_atp_problem format (K []) [(factsN, fact_lines)]
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   310
      in
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   311
        File.append file_order_path (base_name ^ "\n");
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   312
        write_lines (ap include_dir name) ss
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   313
      end
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   314
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   315
    fun select_facts_for_fact facts fact =
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   316
      let
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   317
        val (hyp_ts, conj_t) = Logic.strip_horn (Thm.prop_of (snd fact))
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   318
        val mepo = Sledgehammer_MePo.mepo_suggested_facts ctxt
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   319
          (Sledgehammer_Commands.default_params thy []) max_facts NONE hyp_ts conj_t facts
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   320
      in
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   321
        map (suffix "\n" o fact_name_of o Thm.get_name_hint o snd) mepo
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   322
      end
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   324
    fun write_problem_files _ _ _ _ [] = ()
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   325
      | write_problem_files _ seen_facts includes [] groups =
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   326
        write_problem_files 1 seen_facts includes includes groups
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   327
      | write_problem_files n seen_facts includes _ ((base_name, []) :: groups) =
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   328
        write_problem_files n seen_facts (includes @ [include_line base_name]) [] groups
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   329
      | write_problem_files n seen_facts includes seen_ss
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   330
          ((base_name, fact_line :: fact_lines) :: groups) =
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   331
        let
61860
2ce3d12015b3 cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
blanchet
parents: 61323
diff changeset
   332
          val (alt, pname, sname, conj) =
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   333
            (case fact_line of
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   334
              Formula ((ident, alt), _, phi, _, _) =>
61860
2ce3d12015b3 cleaner generation of metainformation in DFG format and TPTP theory exporter for Sledgehammer
blanchet
parents: 61323
diff changeset
   335
              (alt, problem_name_of base_name n (encode_meta alt),
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   336
               suggestion_name_of base_name n (encode_meta alt),
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   337
               Formula ((ident, alt), Conjecture, phi, NONE, [])))
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   338
          val fact = the (Symtab.lookup fact_tab alt)
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   339
          val fact_s = tptp_string_of_line format fact_line
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   340
        in
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   341
          (if should_generate_problem thy base_name fact_line then
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   342
             let
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51652
diff changeset
   343
               val conj_s = tptp_string_of_line format conj
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   344
             in
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   345
               File.append problem_order_path (pname ^ "\n");
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   346
               write_lines (ap problem_dir pname) (seen_ss @ [conj_s]);
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   347
               write_lines (ap suggestion_dir sname) (select_facts_for_fact facts fact)
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   348
             end
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   349
           else
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   350
             ();
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   351
           write_problem_files (n + 1) (fact :: seen_facts) includes (seen_ss @ [fact_s])
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   352
             ((base_name, fact_lines) :: groups))
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   353
        end
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   354
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   355
    val _ = write_prelude ()
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   356
    val _ = app write_include_file groups
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 60924
diff changeset
   357
    val _ = write_problem_files 1 [] [include_line prelude_base_name] [] groups
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   358
  in () end
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51633
diff changeset
   359
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
   360
end;