src/HOL/TPTP/MaSh_Import.thy
author blanchet
Wed, 11 Jul 2012 21:43:19 +0200
changeset 49260 854a47677335
parent 49254 0016290f904c
child 49265 1065c307fafe
permissions -rw-r--r--
generate ATP dependencies
     1 (*  Title:      HOL/TPTP/MaSh_Import.thy
     2     Author:     Jasmin Blanchette, TU Muenchen
     3 *)
     4 
     5 header {* MaSh Importer *}
     6 
     7 theory MaSh_Import
     8 imports MaSh_Export
     9 uses "mash_import.ML"
    10 begin
    11 
    12 sledgehammer_params
    13   [max_relevant = 40, strict, dont_slice, type_enc = poly_guards??,
    14    lam_trans = combs_and_lifting, timeout = 5, dont_preplay, minimize]
    15 
    16 declare [[sledgehammer_instantiate_inducts]]
    17 
    18 ML {*
    19 open MaSh_Import
    20 *}
    21 
    22 ML {*
    23 val do_it = false (* switch to "true" to generate the files *);
    24 val thy = @{theory List}
    25 *}
    26 
    27 ML {*
    28 if do_it then import_and_evaluate_mash_suggestions @{context} thy "/tmp/mash_suggestions_list"
    29 else ()
    30 *}
    31 
    32 end