src/HOL/TPTP/ATP_Theory_Export.thy
author wenzelm
Tue, 24 Sep 2024 21:31:20 +0200
changeset 80946 b76f64d7d493
parent 74902 ece4f07ebb04
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
46321
484dc68c8c89 renamed theory exporter
blanchet
parents: 45305
diff changeset
     1
(*  Title:      HOL/TPTP/ATP_Theory_Export.thy
484dc68c8c89 renamed theory exporter
blanchet
parents: 45305
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
484dc68c8c89 renamed theory exporter
blanchet
parents: 45305
diff changeset
     3
*)
484dc68c8c89 renamed theory exporter
blanchet
parents: 45305
diff changeset
     4
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61323
diff changeset
     5
section \<open>ATP Theory Exporter\<close>
46321
484dc68c8c89 renamed theory exporter
blanchet
parents: 45305
diff changeset
     6
484dc68c8c89 renamed theory exporter
blanchet
parents: 45305
diff changeset
     7
theory ATP_Theory_Export
51648
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
     8
imports Complex_Main
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
     9
begin
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    10
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69597
diff changeset
    11
ML_file \<open>atp_theory_export.ML\<close>
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 48375
diff changeset
    12
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61323
diff changeset
    13
ML \<open>
48375
blanchet
parents: 48302
diff changeset
    14
open ATP_Problem
blanchet
parents: 48302
diff changeset
    15
open ATP_Theory_Export
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61323
diff changeset
    16
\<close>
45305
3e09961326ce also export DFG formats
blanchet
parents: 44402
diff changeset
    17
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61323
diff changeset
    18
ML \<open>
51648
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    19
val do_it = false (* switch to "true" to generate the files *)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63167
diff changeset
    20
val ctxt = \<^context>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 63167
diff changeset
    21
val thy = \<^theory>\<open>Complex_Main\<close>
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 58889
diff changeset
    22
val infer_policy = (* Unchecked_Inferences *) No_Inferences
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61323
diff changeset
    23
\<close>
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 58889
diff changeset
    24
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61323
diff changeset
    25
ML \<open>
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 58889
diff changeset
    26
if do_it then
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 58889
diff changeset
    27
  "/tmp/isa_filter"
72591
56514a42eee8 Updated ML in forgotten in previous commit
desharna
parents: 69605
diff changeset
    28
  |> generate_casc_lbt_isa_files_for_theory ctxt thy
74902
ece4f07ebb04 fixed HOL-TPTP
desharna
parents: 72591
diff changeset
    29
     (THF (Polymorphic, {with_ite = false, with_let = false}, THF_Without_Choice)) infer_policy
ece4f07ebb04 fixed HOL-TPTP
desharna
parents: 72591
diff changeset
    30
     "poly_native_higher"
61323
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 58889
diff changeset
    31
else
99b3a17a7eab extended theory exporter to also export MePo-selected facts
blanchet
parents: 58889
diff changeset
    32
  ()
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61323
diff changeset
    33
\<close>
51648
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    34
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61323
diff changeset
    35
ML \<open>
51648
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    36
if do_it then
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    37
  "/tmp/axs_tc_native.dfg"
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    38
  |> generate_atp_inference_file_for_theory ctxt thy (DFG Polymorphic)
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    39
         infer_policy "tc_native"
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    40
else
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    41
  ()
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61323
diff changeset
    42
\<close>
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    43
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61323
diff changeset
    44
ML \<open>
51648
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    45
if do_it then
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    46
  "/tmp/infs_poly_guards_query_query.tptp"
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    47
  |> generate_atp_inference_file_for_theory ctxt thy FOF infer_policy
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    48
         "poly_guards??"
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    49
else
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    50
  ()
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61323
diff changeset
    51
\<close>
48301
e5c5037a3104 started implementing MaSh client-side I/O
blanchet
parents: 48239
diff changeset
    52
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61323
diff changeset
    53
ML \<open>
51648
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    54
if do_it then
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    55
  "/tmp/infs_poly_tags_query_query.tptp"
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    56
  |> generate_atp_inference_file_for_theory ctxt thy FOF infer_policy
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    57
         "poly_tags??"
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    58
else
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    59
  ()
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61323
diff changeset
    60
\<close>
51646
005b7682178b work on CASC LTB ISA exporter
blanchet
parents: 51645
diff changeset
    61
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61323
diff changeset
    62
ML \<open>
51648
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    63
if do_it then
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    64
  "/tmp/casc_ltb_isa"
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    65
  |> generate_casc_lbt_isa_files_for_theory ctxt thy FOF infer_policy
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    66
         "poly_tags??"
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    67
else
3e09226c3378 reverted accidental changes to theory file + updated wrt ML file
blanchet
parents: 51646
diff changeset
    68
  ()
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61323
diff changeset
    69
\<close>
43400
3d03f4472883 use more appropriate type systems for ATP exporter
blanchet
parents: 43223
diff changeset
    70
42602
a2db47fa015e added TPTP exporter facility -- useful to do experiments with machine learning
blanchet
parents:
diff changeset
    71
end