src/HOL/TPTP/MaSh_Export.thy
author blanchet
Fri, 30 May 2014 12:27:51 +0200
changeset 58464 5f69b8c3af5a
parent 58463 426ab5fabcae
child 58492 f591096a9c94
permissions -rw-r--r--
more work on exporter
     1 (*  Title:      HOL/TPTP/MaSh_Export.thy
     2     Author:     Jasmin Blanchette, TU Muenchen
     3 *)
     4 
     5 header {* MaSh Exporter *}
     6 
     7 theory MaSh_Export
     8 imports Main
     9 begin
    10 
    11 ML_file "mash_export.ML"
    12 
    13 sledgehammer_params
    14   [provers = spass, max_facts = 32, strict, dont_slice, type_enc = mono_native,
    15    lam_trans = lifting, timeout = 2, dont_preplay, minimize]
    16 
    17 declare [[sledgehammer_instantiate_inducts = false]]
    18 
    19 hide_fact (open) HOL.ext
    20 
    21 ML {*
    22 Multithreading.max_threads_value ()
    23 *}
    24 
    25 ML {*
    26 open MaSh_Export
    27 *}
    28 
    29 ML {*
    30 val do_it = false (* switch to "true" to generate the files *)
    31 val thys = [@{theory List}]
    32 val params as {provers, ...} = Sledgehammer_Commands.default_params @{theory} []
    33 val prover = hd provers
    34 val range = (1, NONE)
    35 val step = 1
    36 val max_suggestions = 1024
    37 val dir = "List"
    38 val prefix = "/tmp/" ^ dir ^ "/"
    39 *}
    40 
    41 ML {*
    42 if do_it then
    43   Isabelle_System.mkdir (Path.explode prefix)
    44 else
    45   ()
    46 *}
    47 
    48 ML {* Options.put_default @{system_option MaSh} "sml_nb" *}
    49 
    50 ML {*
    51 if do_it then
    52   generate_mash_suggestions @{context} params (range, step) thys max_suggestions
    53     (prefix ^ "mash_sml_nb_suggestions")
    54 else
    55   ()
    56 *}
    57 
    58 ML {* Options.put_default @{system_option MaSh} "sml_knn" *}
    59 
    60 ML {*
    61 if do_it then
    62   generate_mash_suggestions @{context} params (range, step) thys max_suggestions
    63     (prefix ^ "mash_sml_knn_suggestions")
    64 else
    65   ()
    66 *}
    67 
    68 ML {* Options.put_default @{system_option MaSh} "py" *}
    69 
    70 ML {*
    71 if do_it then
    72   generate_mash_suggestions @{context} params (range, step) thys max_suggestions
    73     (prefix ^ "mash_py_suggestions")
    74 else
    75   ()
    76 *}
    77 
    78 ML {*
    79 if do_it then
    80   generate_accessibility @{context} thys (prefix ^ "mash_accessibility")
    81 else
    82   ()
    83 *}
    84 
    85 ML {*
    86 if do_it then
    87   generate_features @{context} thys (prefix ^ "mash_features")
    88 else
    89   ()
    90 *}
    91 
    92 ML {*
    93 if do_it then
    94   generate_isar_dependencies @{context} range thys (prefix ^ "mash_dependencies")
    95 else
    96   ()
    97 *}
    98 
    99 ML {*
   100 if do_it then
   101   generate_isar_commands @{context} prover (range, step) thys max_suggestions
   102     (prefix ^ "mash_commands")
   103 else
   104   ()
   105 *}
   106 
   107 ML {*
   108 if do_it then
   109   generate_mepo_suggestions @{context} params (range, step) thys max_suggestions
   110     (prefix ^ "mepo_suggestions")
   111 else
   112   ()
   113 *}
   114 
   115 ML {*
   116 if do_it then
   117   generate_mesh_suggestions max_suggestions (prefix ^ "mash_suggestions")
   118     (prefix ^ "mepo_suggestions") (prefix ^ "mesh_suggestions")
   119 else
   120   ()
   121 *}
   122 
   123 ML {*
   124 if do_it then
   125   generate_prover_dependencies @{context} params range thys (prefix ^ "mash_prover_dependencies")
   126 else
   127   ()
   128 *}
   129 
   130 ML {*
   131 if do_it then
   132   generate_prover_commands @{context} params (range, step) thys max_suggestions
   133     (prefix ^ "mash_prover_commands")
   134 else
   135   ()
   136 *}
   137 
   138 ML {*
   139 if do_it then
   140   generate_mesh_suggestions max_suggestions (prefix ^ "mash_prover_suggestions")
   141     (prefix ^ "mepo_suggestions") (prefix ^ "mesh_prover_suggestions")
   142 else
   143   ()
   144 *}
   145 
   146 end