| author | wenzelm | 
| Sat, 27 Jul 2013 16:44:58 +0200 | |
| changeset 52734 | 077149654ab4 | 
| parent 47687 | bfbd2d0bb348 | 
| child 55595 | 2e2e9bc7c4c6 | 
| permissions | -rw-r--r-- | 
| 47512 | 1 | (* Title: HOL/TPTP/TPTP_Test.thy | 
| 47518 | 2 | Author: Nik Sultana, Cambridge University Computer Laboratory | 
| 47512 | 3 | |
| 47558 
55b42f9af99d
phase out "$TPTP_PROBLEMS_PATH"; prefer "$TPTP" for consistency with CASC setup
 blanchet parents: 
47547diff
changeset | 4 | Some test support for the TPTP interface. Some of the tests rely on the Isabelle | 
| 
55b42f9af99d
phase out "$TPTP_PROBLEMS_PATH"; prefer "$TPTP" for consistency with CASC setup
 blanchet parents: 
47547diff
changeset | 5 | environment variable $TPTP, which should point to the TPTP-vX.Y.Z directory. | 
| 47512 | 6 | *) | 
| 7 | ||
| 8 | theory TPTP_Test | |
| 9 | imports TPTP_Parser | |
| 10 | begin | |
| 11 | ||
| 12 | ML {*
 | |
| 47518 | 13 |   val tptp_test_out = Attrib.setup_config_string @{binding "tptp_test_out"} (K "")
 | 
| 14 |   val tptp_test_all = Attrib.setup_config_bool @{binding "tptp_test_all"} (K false)
 | |
| 15 | val tptp_test_timeout = | |
| 16 |     Attrib.setup_config_int @{binding "tptp_test_timeout"} (K 5)
 | |
| 17 | fun test_all ctxt = Config.get ctxt tptp_test_all | |
| 18 | fun get_timeout ctxt = Config.get ctxt tptp_test_timeout | |
| 47512 | 19 | fun S x y z = x z (y z) | 
| 20 | *} | |
| 21 | ||
| 22 | section "Parser tests" | |
| 23 | ||
| 24 | ML {*
 | |
| 25 | fun payload_of (TPTP_Syntax.Annotated_Formula (_, _, _, _, fmla, _)) = fmla | |
| 26 | val payloads_of = map payload_of | |
| 27 | *} | |
| 28 | ||
| 29 | ||
| 30 | section "Source problems" | |
| 31 | ML {*
 | |
| 32 | (*problem source*) | |
| 33 | val tptp_probs_dir = | |
| 47558 
55b42f9af99d
phase out "$TPTP_PROBLEMS_PATH"; prefer "$TPTP" for consistency with CASC setup
 blanchet parents: 
47547diff
changeset | 34 | Path.explode "$TPTP/Problems" | 
| 47512 | 35 | |> Path.expand; | 
| 36 | ||
| 37 | (*list of files to under test*) | |
| 47687 | 38 | fun test_files () = TPTP_Syntax.get_file_list tptp_probs_dir; | 
| 47512 | 39 | *} | 
| 40 | ||
| 41 | ||
| 42 | section "Supporting test functions" | |
| 43 | ML {*
 | |
| 44 | fun report ctxt str = | |
| 45 | let | |
| 47518 | 46 | val tptp_test_out = Config.get ctxt tptp_test_out | 
| 47512 | 47 | in | 
| 47518 | 48 | if tptp_test_out = "" then warning str | 
| 47512 | 49 | else | 
| 50 | let | |
| 47518 | 51 | val out_stream = TextIO.openAppend tptp_test_out | 
| 47512 | 52 | in (TextIO.output (out_stream, str ^ "\n"); | 
| 53 | TextIO.flushOut out_stream; | |
| 54 | TextIO.closeOut out_stream) | |
| 55 | end | |
| 56 | end | |
| 57 | ||
| 58 | fun test_fn ctxt f msg default_val file_name = | |
| 59 | let | |
| 60 | val _ = TPTP_Syntax.debug tracing (msg ^ " " ^ Path.print file_name) | |
| 61 | in | |
| 62 | (f file_name; ()) | |
| 63 | (*otherwise report exceptions as warnings*) | |
| 64 | handle exn => | |
| 65 | if Exn.is_interrupt exn then | |
| 66 | reraise exn | |
| 67 | else | |
| 68 | (report ctxt (msg ^ " test: file " ^ Path.print file_name ^ | |
| 69 | " raised exception: " ^ ML_Compiler.exn_message exn); | |
| 70 | default_val) | |
| 71 | end | |
| 72 | ||
| 73 | fun timed_test ctxt f = | |
| 74 | let | |
| 75 | fun f' x = (f x; ()) | |
| 76 | val time = | |
| 47687 | 77 | Timing.timing (List.app f') (test_files ()) | 
| 47512 | 78 | |> fst | 
| 79 | val duration = | |
| 80 | #elapsed time | |
| 81 | |> Time.toSeconds | |
| 82 | |> Real.fromLargeInt | |
| 83 | val average = | |
| 84 | (StringCvt.FIX (SOME 3), | |
| 47687 | 85 | (duration / Real.fromInt (length (test_files ())))) | 
| 47512 | 86 | |-> Real.fmt | 
| 87 | in | |
| 88 |       report ctxt ("Test timing: " ^ Timing.message time ^ "\n(about " ^ average ^
 | |
| 89 | "s per problem)") | |
| 90 | end | |
| 91 | *} | |
| 92 | ||
| 93 | ||
| 94 | ML {*
 | |
| 95 | fun situate file_name = Path.append tptp_probs_dir (Path.explode file_name); | |
| 47547 | 96 | |
| 47512 | 97 | fun parser_test ctxt = (*FIXME argument order*) | 
| 98 | test_fn ctxt | |
| 99 | (fn file_name => | |
| 100 | Path.implode file_name | |
| 101 | |> (fn file => | |
| 102 | ((*report ctxt file; this is if you want the filename in the log*) | |
| 103 | TPTP_Parser.parse_file file))) | |
| 104 | "parser" | |
| 105 | () | |
| 47547 | 106 | |
| 107 | fun parse_timed file = | |
| 108 | Timing.timing TPTP_Parser.parse_file (Path.implode file) | |
| 47512 | 109 | *} | 
| 110 | ||
| 111 | end |