src/HOL/TPTP/MaSh_Eval.thy
author paulson <lp15@cam.ac.uk>
Tue, 31 Mar 2015 16:48:48 +0100
changeset 59865 8a20dd967385
parent 58889 5b7a9633cfa8
child 62925 f1bdf10f95d8
permissions -rw-r--r--
rationalised and generalised some theorems concerning abs and x^2.

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

section {* MaSh Evaluation Driver *}

theory MaSh_Eval
imports MaSh_Export_Base
begin

ML_file "mash_eval.ML"

sledgehammer_params
  [provers = e, max_facts = 64, strict, dont_slice, type_enc = poly_guards??,
   lam_trans = combs, timeout = 30, dont_preplay, minimize]

ML {*
Multithreading.max_threads_value ()
*}

ML {*
open MaSh_Eval
*}

ML {*
val params = Sledgehammer_Commands.default_params @{theory} []
val prob_dir = prefix ^ "mash_problems"
*}

ML {*
if do_it then
  Isabelle_System.mkdir (Path.explode prob_dir)
else
  ()
*}

ML {*
if do_it then
  evaluate_mash_suggestions @{context} params range (SOME prob_dir)
    [prefix ^ "mepo_suggestions",
     prefix ^ "mash_suggestions",
     prefix ^ "mash_prover_suggestions",
     prefix ^ "mesh_suggestions",
     prefix ^ "mesh_prover_suggestions"]
    (prefix ^ "mash_eval")
else
  ()
*}

end