src/HOL/TPTP/MaSh_Export.thy
author blanchet
Wed, 18 Jul 2012 08:44:03 +0200
changeset 49316 e5c5037a3104
parent 49314 5e5c6616f0fe
child 49319 50e64af9c829
permissions -rw-r--r--
started implementing MaSh client-side I/O
     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 (* ### Complex_Main *) "~~/src/HOL/Sledgehammer2d"
     9 uses "mash_export.ML"
    10 begin
    11 
    12 sledgehammer_params
    13   [provers = e, max_relevant = 40, strict, dont_slice, type_enc = poly_guards??,
    14    lam_trans = combs_and_lifting, timeout = 5, dont_preplay, minimize]
    15 
    16 ML {*
    17 open Sledgehammer_Filter_MaSh
    18 open MaSh_Export
    19 *}
    20 
    21 ML {*
    22 val do_it = false (* switch to "true" to generate the files *);
    23 val thy = @{theory List};
    24 val params = Sledgehammer_Isar.default_params @{context} []
    25 *}
    26 
    27 ML {*
    28 if do_it then
    29   generate_accessibility @{context} thy false "/tmp/mash_accessibility"
    30 else
    31   ()
    32 *}
    33 
    34 ML {*
    35 if do_it then
    36   generate_features @{context} thy false "/tmp/mash_features"
    37 else
    38   ()
    39 *}
    40 
    41 ML {*
    42 if do_it then
    43   generate_isa_dependencies @{context} thy false "/tmp/mash_isa_dependencies"
    44 else
    45   ()
    46 *}
    47 
    48 ML {*
    49 if do_it then
    50   generate_atp_dependencies @{context} params thy false "/tmp/mash_atp_dependencies"
    51 else
    52   ()
    53 *}
    54 
    55 ML {*
    56 if do_it then
    57   generate_commands @{context} thy "/tmp/mash_commands"
    58 else
    59   ()
    60 *}
    61 
    62 ML {*
    63 if do_it then
    64   generate_iter_suggestions @{context} params thy 500 "/tmp/mash_iter_suggestions"
    65 else
    66   ()
    67 *}
    68 
    69 end