src/HOL/TPTP/MaSh_Export.thy
author blanchet
Wed, 18 Jul 2012 08:44:04 +0200
changeset 49319 50e64af9c829
parent 49316 e5c5037a3104
child 49348 2250197977dc
permissions -rw-r--r--
more work on MaSh
     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 MaSh_Export
    18 *}
    19 
    20 ML {*
    21 val do_it = false (* switch to "true" to generate the files *);
    22 val thy = @{theory List};
    23 val params = Sledgehammer_Isar.default_params @{context} []
    24 *}
    25 
    26 ML {*
    27 if do_it then
    28   generate_accessibility thy false "/tmp/mash_accessibility"
    29 else
    30   ()
    31 *}
    32 
    33 ML {*
    34 if do_it then
    35   generate_features @{context} thy false "/tmp/mash_features"
    36 else
    37   ()
    38 *}
    39 
    40 ML {*
    41 if do_it then
    42   generate_isa_dependencies thy false "/tmp/mash_isa_dependencies"
    43 else
    44   ()
    45 *}
    46 
    47 ML {*
    48 if do_it then
    49   generate_atp_dependencies @{context} params thy false "/tmp/mash_atp_dependencies"
    50 else
    51   ()
    52 *}
    53 
    54 ML {*
    55 if do_it then
    56   generate_commands @{context} thy "/tmp/mash_commands"
    57 else
    58   ()
    59 *}
    60 
    61 ML {*
    62 if do_it then
    63   generate_iter_suggestions @{context} params thy 500 "/tmp/mash_iter_suggestions"
    64 else
    65   ()
    66 *}
    67 
    68 end