src/HOL/TPTP/MaSh_Export.thy
author blanchet
Thu, 06 Dec 2012 16:49:48 +0100
changeset 51426 c9023d78d1a6
parent 51365 136d5318b1fe
child 51449 960a3429615c
permissions -rw-r--r--
export ATP and Isar commands separately
     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_relevant = 32, strict, dont_slice, type_enc = mono_native,
    15    lam_trans = combs_and_lifting, timeout = 2, dont_preplay, minimize]
    16 
    17 ML {*
    18 open MaSh_Export
    19 *}
    20 
    21 ML {*
    22 val do_it = false (* switch to "true" to generate the files *)
    23 val thys = [@{theory List}]
    24 val params as {provers, ...} = Sledgehammer_Isar.default_params @{context} []
    25 val prover = hd provers
    26 val dir = space_implode "__" (map Context.theory_name thys)
    27 val prefix = "/tmp/" ^ dir ^ "/"
    28 *}
    29 
    30 ML {*
    31 if do_it then
    32   Isabelle_System.mkdir (Path.explode prefix)
    33 else
    34   ()
    35 *}
    36 
    37 ML {*
    38 if do_it then
    39   generate_accessibility @{context} thys false (prefix ^ "mash_accessibility")
    40 else
    41   ()
    42 *}
    43 
    44 ML {*
    45 if do_it then
    46   generate_features @{context} prover thys false (prefix ^ "mash_features")
    47 else
    48   ()
    49 *}
    50 
    51 ML {*
    52 if do_it then
    53   generate_isar_dependencies @{context} thys false (prefix ^ "mash_dependencies")
    54 else
    55   ()
    56 *}
    57 
    58 ML {*
    59 if do_it then
    60   generate_isar_commands @{context} prover thys (prefix ^ "mash_commands")
    61 else
    62   ()
    63 *}
    64 
    65 ML {*
    66 if do_it then
    67   generate_mepo_suggestions @{context} params thys 1024 (prefix ^ "mash_mepo_suggestions")
    68 else
    69   ()
    70 *}
    71 
    72 ML {*
    73 if do_it then
    74   generate_atp_dependencies @{context} params thys false (prefix ^ "mash_atp_dependencies")
    75 else
    76   ()
    77 *}
    78 
    79 ML {*
    80 if do_it then
    81   generate_atp_commands @{context} params thys (prefix ^ "mash_atp_commands")
    82 else
    83   ()
    84 *}
    85 
    86 end