src/HOL/TPTP/MaSh_Export.thy
author wenzelm
Wed, 22 Aug 2012 22:55:41 +0200
changeset 49906 c0eafbd55de3
parent 49546 7da5d3b8aef4
child 51352 68555697f37e
permissions -rw-r--r--
prefer ML_file over old uses;
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
begin
blanchet@49249
    10
wenzelm@49906
    11
ML_file "mash_export.ML"
wenzelm@49906
    12
blanchet@49260
    13
sledgehammer_params
blanchet@49260
    14
  [provers = e, max_relevant = 40, strict, dont_slice, type_enc = poly_guards??,
blanchet@49447
    15
   lam_trans = combs_and_lifting, timeout = 1, dont_preplay, minimize]
blanchet@49254
    16
blanchet@49249
    17
ML {*
blanchet@49249
    18
open MaSh_Export
blanchet@49249
    19
*}
blanchet@49249
    20
blanchet@49249
    21
ML {*
blanchet@49348
    22
val do_it = false (* switch to "true" to generate the files *)
blanchet@49348
    23
val thy = @{theory List}
blanchet@49348
    24
val params as {provers, ...} = Sledgehammer_Isar.default_params @{context} []
blanchet@49348
    25
val prover = hd provers
blanchet@49261
    26
*}
blanchet@49546
    27
blanchet@49261
    28
ML {*
blanchet@49260
    29
if do_it then
blanchet@49546
    30
  generate_accessibility @{context} thy false "/tmp/mash_accessibility"
blanchet@49260
    31
else
blanchet@49260
    32
  ()
blanchet@49249
    33
*}
blanchet@49249
    34
blanchet@49249
    35
ML {*
blanchet@49260
    36
if do_it then
blanchet@49348
    37
  generate_features @{context} prover thy false "/tmp/mash_features"
blanchet@49260
    38
else
blanchet@49260
    39
  ()
blanchet@49249
    40
*}
blanchet@49249
    41
blanchet@49249
    42
ML {*
blanchet@49260
    43
if do_it then
blanchet@49546
    44
  generate_isar_dependencies @{context} thy false "/tmp/mash_dependencies"
blanchet@49348
    45
else
blanchet@49348
    46
  ()
blanchet@49348
    47
*}
blanchet@49348
    48
blanchet@49348
    49
ML {*
blanchet@49348
    50
if do_it then
blanchet@49544
    51
  generate_commands @{context} params thy "/tmp/mash_commands"
blanchet@49348
    52
else
blanchet@49348
    53
  ()
blanchet@49348
    54
*}
blanchet@49348
    55
blanchet@49348
    56
ML {*
blanchet@49348
    57
if do_it then
blanchet@49394
    58
  generate_mepo_suggestions @{context} params thy 500 "/tmp/mash_mepo_suggestions"
blanchet@49260
    59
else
blanchet@49260
    60
  ()
blanchet@49249
    61
*}
blanchet@49249
    62
blanchet@49249
    63
ML {*
blanchet@49260
    64
if do_it then
blanchet@49266
    65
  generate_atp_dependencies @{context} params thy false "/tmp/mash_atp_dependencies"
blanchet@49260
    66
else
blanchet@49260
    67
  ()
blanchet@49260
    68
*}
blanchet@49260
    69
blanchet@49249
    70
end