src/HOL/TPTP/MaSh_Export.thy
author blanchet
Fri, 31 Jan 2014 10:23:32 +0100
changeset 56540 7a538e58b64e
parent 56059 42c209a6c225
child 56550 11dd3d1da83b
permissions -rw-r--r--
tuned ML file name
     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 @{context} []
    33 val prover = hd provers
    34 val range = (1, NONE)
    35 val step = 1
    36 val linearize = false
    37 val max_suggestions = 1024
    38 val dir = "List"
    39 val prefix = "/tmp/" ^ dir ^ "/"
    40 *}
    41 
    42 ML {*
    43 if do_it then
    44   Isabelle_System.mkdir (Path.explode prefix)
    45 else
    46   ()
    47 *}
    48 
    49 ML {*
    50 if do_it then
    51   generate_accessibility @{context} thys linearize
    52       (prefix ^ "mash_accessibility")
    53 else
    54   ()
    55 *}
    56 
    57 ML {*
    58 if do_it then
    59   generate_features @{context} prover thys (prefix ^ "mash_features")
    60 else
    61   ()
    62 *}
    63 
    64 ML {*
    65 if do_it then
    66   generate_isar_dependencies @{context} range thys linearize
    67       (prefix ^ "mash_dependencies")
    68 else
    69   ()
    70 *}
    71 
    72 ML {*
    73 if do_it then
    74   generate_isar_commands @{context} prover (range, step) thys linearize
    75       max_suggestions (prefix ^ "mash_commands")
    76 else
    77   ()
    78 *}
    79 
    80 ML {*
    81 if do_it then
    82   generate_mepo_suggestions @{context} params (range, step) thys linearize
    83       max_suggestions (prefix ^ "mepo_suggestions")
    84 else
    85   ()
    86 *}
    87 
    88 ML {*
    89 if do_it then
    90   generate_mesh_suggestions max_suggestions (prefix ^ "mash_suggestions")
    91       (prefix ^ "mepo_suggestions") (prefix ^ "mesh_suggestions")
    92 else
    93   ()
    94 *}
    95 
    96 ML {*
    97 if do_it then
    98   generate_prover_dependencies @{context} params range thys linearize
    99       (prefix ^ "mash_prover_dependencies")
   100 else
   101   ()
   102 *}
   103 
   104 ML {*
   105 if do_it then
   106   generate_prover_commands @{context} params (range, step) thys linearize
   107       max_suggestions (prefix ^ "mash_prover_commands")
   108 else
   109   ()
   110 *}
   111 
   112 ML {*
   113 if do_it then
   114   generate_mesh_suggestions max_suggestions (prefix ^ "mash_prover_suggestions")
   115       (prefix ^ "mepo_suggestions") (prefix ^ "mesh_prover_suggestions")
   116 else
   117   ()
   118 *}
   119 
   120 end