src/HOL/TPTP/MaSh_Export.thy
author blanchet
Wed, 11 Jul 2012 21:43:19 +0200
changeset 49261 fb11c09d7729
parent 49260 854a47677335
child 49265 1065c307fafe
permissions -rw-r--r--
add Isabelle dependencies to tweak relevance filter
     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 ATP_Theory_Export
     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 = true (* switch to "true" to generate the files *);
    22 val thy = @{theory Nat}
    23 *}
    24 
    25 ML {*
    26 if do_it then
    27   generate_atp_dependencies @{context} thy false "/tmp/mash_atp_dependencies"
    28 else
    29   ()
    30 *}
    31 
    32 ML {*
    33 if do_it then
    34   generate_features thy false "/tmp/mash_features"
    35 else
    36   ()
    37 *}
    38 
    39 ML {*
    40 if do_it then
    41   generate_accessibility thy false "/tmp/mash_accessibility"
    42 else
    43   ()
    44 *}
    45 
    46 ML {*
    47 if do_it then
    48   generate_isa_dependencies thy false "/tmp/mash_isa_dependencies"
    49 else
    50   ()
    51 *}
    52 
    53 ML {*
    54 if do_it then
    55   generate_atp_dependencies @{context} thy false "/tmp/mash_atp_dependencies"
    56 else
    57   ()
    58 *}
    59 
    60 ML {*
    61 if do_it then
    62   generate_commands thy "/tmp/mash_commands"
    63 else
    64   ()
    65 *}
    66 
    67 ML {*
    68 if do_it then
    69   generate_meng_paulson_suggestions @{context} thy 500 "/tmp/mash_meng_paulson_suggestions"
    70 else
    71   ()
    72 *}
    73 
    74 end