src/HOL/TPTP/MaSh_Export.thy
author blanchet
Mon, 23 Jul 2012 15:32:30 +0200
changeset 49447 60759d07df24
parent 49394 2b5ad61e2ccc
child 49544 716ec3458b1d
permissions -rw-r--r--
took out CVC3 again -- there seems to be issues with the server version of CVC3 + minor tweaks
blanchet@49249
     1
(*  Title:      HOL/TPTP/MaSh_Export.thy
blanchet@49249
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@49249
     3
*)
blanchet@49249
     4
blanchet@49249
     5
header {* MaSh Exporter *}
blanchet@49249
     6
blanchet@49249
     7
theory MaSh_Export
blanchet@49348
     8
imports Complex_Main
blanchet@49249
     9
uses "mash_export.ML"
blanchet@49249
    10
begin
blanchet@49249
    11
blanchet@49260
    12
sledgehammer_params
blanchet@49260
    13
  [provers = e, max_relevant = 40, strict, dont_slice, type_enc = poly_guards??,
blanchet@49447
    14
   lam_trans = combs_and_lifting, timeout = 1, dont_preplay, minimize]
blanchet@49254
    15
blanchet@49249
    16
ML {*
blanchet@49249
    17
open MaSh_Export
blanchet@49249
    18
*}
blanchet@49249
    19
blanchet@49249
    20
ML {*
blanchet@49348
    21
val do_it = false (* switch to "true" to generate the files *)
blanchet@49348
    22
val thy = @{theory List}
blanchet@49348
    23
val params as {provers, ...} = Sledgehammer_Isar.default_params @{context} []
blanchet@49348
    24
val prover = hd provers
blanchet@49261
    25
*}
blanchet@49261
    26
blanchet@49261
    27
ML {*
blanchet@49260
    28
if do_it then
blanchet@49319
    29
  generate_accessibility thy false "/tmp/mash_accessibility"
blanchet@49260
    30
else
blanchet@49260
    31
  ()
blanchet@49249
    32
*}
blanchet@49249
    33
blanchet@49249
    34
ML {*
blanchet@49260
    35
if do_it then
blanchet@49348
    36
  generate_features @{context} prover thy false "/tmp/mash_features"
blanchet@49260
    37
else
blanchet@49260
    38
  ()
blanchet@49249
    39
*}
blanchet@49249
    40
blanchet@49249
    41
ML {*
blanchet@49260
    42
if do_it then
blanchet@49348
    43
  generate_isar_dependencies thy false "/tmp/mash_dependencies"
blanchet@49348
    44
else
blanchet@49348
    45
  ()
blanchet@49348
    46
*}
blanchet@49348
    47
blanchet@49348
    48
ML {*
blanchet@49348
    49
if do_it then
blanchet@49348
    50
  generate_commands @{context} prover thy "/tmp/mash_commands"
blanchet@49348
    51
else
blanchet@49348
    52
  ()
blanchet@49348
    53
*}
blanchet@49348
    54
blanchet@49348
    55
ML {*
blanchet@49348
    56
if do_it then
blanchet@49394
    57
  generate_mepo_suggestions @{context} params thy 500 "/tmp/mash_mepo_suggestions"
blanchet@49260
    58
else
blanchet@49260
    59
  ()
blanchet@49249
    60
*}
blanchet@49249
    61
blanchet@49249
    62
ML {*
blanchet@49260
    63
if do_it then
blanchet@49266
    64
  generate_atp_dependencies @{context} params thy false "/tmp/mash_atp_dependencies"
blanchet@49260
    65
else
blanchet@49260
    66
  ()
blanchet@49260
    67
*}
blanchet@49260
    68
blanchet@49249
    69
end