blanchet@49249: (* Title: HOL/TPTP/MaSh_Export.thy blanchet@49249: Author: Jasmin Blanchette, TU Muenchen blanchet@49249: *) blanchet@49249: blanchet@49249: header {* MaSh Exporter *} blanchet@49249: blanchet@49249: theory MaSh_Export blanchet@49348: imports Complex_Main blanchet@49249: uses "mash_export.ML" blanchet@49249: begin blanchet@49249: blanchet@49260: sledgehammer_params blanchet@49260: [provers = e, max_relevant = 40, strict, dont_slice, type_enc = poly_guards??, blanchet@49447: lam_trans = combs_and_lifting, timeout = 1, dont_preplay, minimize] blanchet@49254: blanchet@49249: ML {* blanchet@49249: open MaSh_Export blanchet@49249: *} blanchet@49249: blanchet@49249: ML {* blanchet@49348: val do_it = false (* switch to "true" to generate the files *) blanchet@49348: val thy = @{theory List} blanchet@49348: val params as {provers, ...} = Sledgehammer_Isar.default_params @{context} [] blanchet@49348: val prover = hd provers blanchet@49261: *} blanchet@49261: blanchet@49261: ML {* blanchet@49260: if do_it then blanchet@49319: generate_accessibility thy false "/tmp/mash_accessibility" blanchet@49260: else blanchet@49260: () blanchet@49249: *} blanchet@49249: blanchet@49249: ML {* blanchet@49260: if do_it then blanchet@49348: generate_features @{context} prover thy false "/tmp/mash_features" blanchet@49260: else blanchet@49260: () blanchet@49249: *} blanchet@49249: blanchet@49249: ML {* blanchet@49260: if do_it then blanchet@49348: generate_isar_dependencies thy false "/tmp/mash_dependencies" blanchet@49348: else blanchet@49348: () blanchet@49348: *} blanchet@49348: blanchet@49348: ML {* blanchet@49348: if do_it then blanchet@49348: generate_commands @{context} prover thy "/tmp/mash_commands" blanchet@49348: else blanchet@49348: () blanchet@49348: *} blanchet@49348: blanchet@49348: ML {* blanchet@49348: if do_it then blanchet@49394: generate_mepo_suggestions @{context} params thy 500 "/tmp/mash_mepo_suggestions" blanchet@49260: else blanchet@49260: () blanchet@49249: *} blanchet@49249: blanchet@49249: ML {* blanchet@49260: if do_it then blanchet@49266: generate_atp_dependencies @{context} params thy false "/tmp/mash_atp_dependencies" blanchet@49260: else blanchet@49260: () blanchet@49260: *} blanchet@49260: blanchet@49249: end