src/HOL/TPTP/MaSh_Export.thy
author blanchet
Tue, 10 Jul 2012 23:36:03 +0200
changeset 49250 40655464a93b
parent 49249 06216c789ac9
child 49254 0016290f904c
permissions -rw-r--r--
MaSh evaluation driver
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@49249
     8
imports ATP_Theory_Export
blanchet@49249
     9
uses "mash_export.ML"
blanchet@49249
    10
begin
blanchet@49249
    11
blanchet@49249
    12
ML {*
blanchet@49249
    13
open MaSh_Export
blanchet@49249
    14
*}
blanchet@49249
    15
blanchet@49249
    16
ML {*
blanchet@49250
    17
val do_it = false (* switch to "true" to generate the files *)
blanchet@49250
    18
val thy = @{theory List}
blanchet@49249
    19
*}
blanchet@49249
    20
blanchet@49249
    21
ML {*
blanchet@49250
    22
if do_it then generate_mash_commands thy "/tmp/mash_commands"
blanchet@49250
    23
else ()
blanchet@49249
    24
*}
blanchet@49249
    25
blanchet@49249
    26
ML {*
blanchet@49250
    27
if do_it then generate_mash_features thy false "/tmp/mash_features"
blanchet@49250
    28
else ()
blanchet@49249
    29
*}
blanchet@49249
    30
blanchet@49249
    31
ML {*
blanchet@49250
    32
if do_it then generate_mash_accessibility thy false "/tmp/mash_accessibility"
blanchet@49250
    33
else ()
blanchet@49249
    34
*}
blanchet@49249
    35
blanchet@49249
    36
ML {*
blanchet@49250
    37
if do_it then generate_mash_dependencies thy false "/tmp/mash_dependencies"
blanchet@49250
    38
else ()
blanchet@49250
    39
*}
blanchet@49250
    40
blanchet@49250
    41
ML {*
blanchet@49250
    42
find_index (curry (op =) 22) [0, 0, 9, 1, 2, 3]
blanchet@49249
    43
*}
blanchet@49249
    44
blanchet@49249
    45
end