src/HOL/TPTP/MaSh_Export.thy
author blanchet
Thu, 26 Jul 2012 10:48:03 +0200
changeset 48529 716ec3458b1d
parent 48432 60759d07df24
child 48531 7da5d3b8aef4
permissions -rw-r--r--
generate fact name in queries again + use ATP dependencies when possible

(*  Title:      HOL/TPTP/MaSh_Export.thy
    Author:     Jasmin Blanchette, TU Muenchen
*)

header {* MaSh Exporter *}

theory MaSh_Export
imports Complex_Main
uses "mash_export.ML"
begin

sledgehammer_params
  [provers = e, max_relevant = 40, strict, dont_slice, type_enc = poly_guards??,
   lam_trans = combs_and_lifting, timeout = 1, dont_preplay, minimize]

ML {*
open MaSh_Export
*}

ML {*
val do_it = false (* switch to "true" to generate the files *)
val thy = @{theory List}
val params as {provers, ...} = Sledgehammer_Isar.default_params @{context} []
val prover = hd provers
*}
ML {*
if do_it then
  generate_accessibility thy false "/tmp/mash_accessibility"
else
  ()
*}

ML {*
if do_it then
  generate_features @{context} prover thy false "/tmp/mash_features"
else
  ()
*}

ML {*
if do_it then
  generate_isar_dependencies thy false "/tmp/mash_dependencies"
else
  ()
*}

ML {*
if do_it then
  generate_commands @{context} params thy "/tmp/mash_commands"
else
  ()
*}

ML {*
if do_it then
  generate_mepo_suggestions @{context} params thy 500 "/tmp/mash_mepo_suggestions"
else
  ()
*}

ML {*
if do_it then
  generate_atp_dependencies @{context} params thy false "/tmp/mash_atp_dependencies"
else
  ()
*}

end