src/HOL/TPTP/TPTP_Interpret_Test.thy
author haftmann
Mon, 26 Oct 2020 11:28:43 +0000
changeset 72508 c89d8e8bd8c7
parent 63167 0909deb8059b
child 72511 460d743010bc
permissions -rw-r--r--
factored out theory Traditional_Syntax
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47512
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
     1
(*  Title:      HOL/TPTP/TPTP_Interpret_Test.thy
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
     2
    Author:     Nik Sultana, Cambridge University Computer Laboratory
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
     3
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
     4
Some tests for the TPTP interface. Some of the tests rely on the Isabelle
47558
55b42f9af99d phase out "$TPTP_PROBLEMS_PATH"; prefer "$TPTP" for consistency with CASC setup
blanchet
parents: 47548
diff changeset
     5
environment variable $TPTP, which should point to the TPTP-vX.Y.Z directory.
47512
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
     6
*)
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
     7
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
     8
theory TPTP_Interpret_Test
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
     9
imports TPTP_Test TPTP_Interpret
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    10
begin
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    11
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    12
section "Interpreter tests"
47518
b2f209258621 more cleaning of tptp tests;
sultana
parents: 47516
diff changeset
    13
47512
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    14
text "Interpret a problem."
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
    15
ML \<open>
47512
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    16
  val (time, ((type_map, const_map, fmlas), thy)) =
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    17
    Timing.timing
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    18
      (TPTP_Interpret.interpret_file
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    19
       false
48829
6ed588c4f963 look in current directory first before looking up includes in the TPTP directory, as required by Geoff
blanchet
parents: 47693
diff changeset
    20
       [Path.explode "$TPTP"]
47693
64023cf4d148 updated test;
sultana
parents: 47558
diff changeset
    21
       (Path.append tptp_probs_dir (Path.explode "LCL/LCL825-1.p"))
47512
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    22
       []
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    23
       [])
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    24
      @{theory}
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
    25
\<close>
47512
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    26
47693
64023cf4d148 updated test;
sultana
parents: 47558
diff changeset
    27
47512
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    28
text "... and display nicely."
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
    29
ML \<open>
47548
60849d8c457d fixed type interpretation;
sultana
parents: 47547
diff changeset
    30
  List.app (Pretty.writeln o (Syntax.pretty_term @{context}) o #3) fmlas;
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
    31
\<close>
47512
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    32
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    33
subsection "Multiple tests"
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    34
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
    35
ML \<open>
47518
b2f209258621 more cleaning of tptp tests;
sultana
parents: 47516
diff changeset
    36
  (*default timeout is 1 min*)
47516
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    37
  fun interpret timeout file thy =
62519
a564458f94db tuned signature -- clarified modules;
wenzelm
parents: 62505
diff changeset
    38
    Timeout.apply (Time.fromSeconds (if timeout = 0 then 60 else timeout))
47516
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    39
     (TPTP_Interpret.interpret_file
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    40
       false
48829
6ed588c4f963 look in current directory first before looking up includes in the TPTP directory, as required by Geoff
blanchet
parents: 47693
diff changeset
    41
       [Path.explode "$TPTP"]
47516
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    42
       file
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    43
       []
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    44
       []) thy
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    45
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    46
  fun interpret_timed timeout file thy =
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    47
    Timing.timing (interpret timeout file) thy
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    48
47512
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    49
  fun interpretation_test timeout ctxt =
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    50
    test_fn ctxt
47518
b2f209258621 more cleaning of tptp tests;
sultana
parents: 47516
diff changeset
    51
     (fn file => interpret timeout file (Proof_Context.theory_of ctxt))
47512
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    52
     "interpreter"
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    53
     ()
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    54
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    55
  fun interpretation_tests timeout ctxt probs =
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    56
    List.app
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    57
     (interpretation_test timeout ctxt)
58412
f65f11f4854c more standard Isabelle/ML operations;
wenzelm
parents: 56303
diff changeset
    58
     (map situate probs)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
    59
\<close>
47512
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    60
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
    61
ML \<open>
47512
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    62
  val some_probs =
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    63
    ["LCL/LCL825-1.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    64
     "ALG/ALG001^5.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    65
     "COM/COM003+2.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    66
     "COM/COM003-1.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    67
     "COM/COM024^5.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    68
     "DAT/DAT017=1.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    69
     "NUM/NUM021^1.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    70
     "NUM/NUM858=1.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    71
     "SYN/SYN000^2.p"]
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    72
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    73
  val take_too_long =
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    74
    ["NLP/NLP562+1.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    75
     "SWV/SWV546-1.010.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    76
     "SWV/SWV567-1.015.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    77
     "LCL/LCL680+1.020.p"]
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    78
47516
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    79
  val timeouts =
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    80
    ["NUM/NUM923^4.p",
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    81
    "NUM/NUM926^4.p",
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    82
    "NUM/NUM925^4.p",
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    83
    "NUM/NUM924^4.p",
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    84
    "CSR/CSR153^3.p",
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    85
    "CSR/CSR151^3.p",
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    86
    "CSR/CSR148^3.p",
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    87
    "CSR/CSR120^3.p",
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    88
    "CSR/CSR150^3.p",
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    89
    "CSR/CSR119^3.p",
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    90
    "CSR/CSR149^3.p"]
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
    91
47512
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    92
  val more_probs =
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    93
    ["GEG/GEG014^1.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    94
     "GEG/GEG009^1.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    95
     "GEG/GEG004^1.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    96
     "GEG/GEG007^1.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    97
     "GEG/GEG016^1.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    98
     "GEG/GEG024=1.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
    99
     "GEG/GEG010^1.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
   100
     "GEG/GEG003^1.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
   101
     "GEG/GEG018^1.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
   102
     "SYN/SYN045^4.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
   103
     "SYN/SYN001^4.001.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
   104
     "SYN/SYN387^4.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
   105
     "SYN/SYN393^4.002.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
   106
     "SYN/SYN978^4.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
   107
     "SYN/SYN044^4.p",
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
   108
     "SYN/SYN393^4.003.p",
47693
64023cf4d148 updated test;
sultana
parents: 47558
diff changeset
   109
     "SYN/SYN389^4.p",
64023cf4d148 updated test;
sultana
parents: 47558
diff changeset
   110
     "ARI/ARI625=1.p",
64023cf4d148 updated test;
sultana
parents: 47558
diff changeset
   111
     "SYO/SYO561_2.p",
64023cf4d148 updated test;
sultana
parents: 47558
diff changeset
   112
     "SYO/SYO562_1.p",
64023cf4d148 updated test;
sultana
parents: 47558
diff changeset
   113
     "SYN/SYN000_2.p"]
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   114
\<close>
47512
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
   115
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   116
ML \<open>
47547
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   117
 interpretation_tests (get_timeout @{context}) @{context}
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   118
   (some_probs @ take_too_long @ timeouts @ more_probs)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   119
\<close>
47547
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   120
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   121
ML \<open>
47547
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   122
  parse_timed (situate "NUM/NUM923^4.p");
47516
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
   123
  interpret_timed 0 (situate "NUM/NUM923^4.p") @{theory}
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   124
\<close>
47512
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
   125
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   126
ML \<open>
47547
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   127
  fun interp_gain timeout thy file =
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   128
    let
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   129
      val t1 = (parse_timed file |> fst)
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   130
      val t2 = (interpret_timed timeout file thy |> fst)
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   131
        handle exn => (*FIXME*)
62505
9e2a65912111 clarified modules;
wenzelm
parents: 62390
diff changeset
   132
          if Exn.is_interrupt exn then Exn.reraise exn
47547
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   133
          else
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   134
            (warning (" test: file " ^ Path.print file ^
56303
4cc3f4db3447 clarified Isabelle/ML bootstrap, such that Execution does not require ML_Compiler;
wenzelm
parents: 55595
diff changeset
   135
             " raised exception: " ^ Runtime.exn_message exn);
47547
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   136
             {gc = Time.zeroTime, cpu = Time.zeroTime, elapsed = Time.zeroTime})
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   137
      val to_real = Time.toReal
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   138
      val diff_elapsed =
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   139
        #elapsed t2 - #elapsed t1
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   140
        |> to_real
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   141
      val elapsed = to_real (#elapsed t2)
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   142
    in
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   143
      (Path.base file, diff_elapsed,
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   144
       diff_elapsed / elapsed,
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   145
       elapsed)
1a5dc8377b5c more tptp testing support functions;
sultana
parents: 47518
diff changeset
   146
    end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   147
\<close>
47516
9c29589c9b31 improved tptp interpretation test thy
sultana
parents: 47512
diff changeset
   148
47512
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
   149
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
   150
subsection "Test against whole TPTP"
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
   151
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
   152
text "Run interpretation over all problems. This works only for logics
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
   153
 for which interpretation is defined (in TPTP_Parser/tptp_interpret.ML)."
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   154
ML \<open>
47518
b2f209258621 more cleaning of tptp tests;
sultana
parents: 47516
diff changeset
   155
  if test_all @{context} then
b2f209258621 more cleaning of tptp tests;
sultana
parents: 47516
diff changeset
   156
    (report @{context} "Interpreting all problems";
55595
2e2e9bc7c4c6 made list of test files a parameter to timed_test in TPTP_Test.thy, and updated dependent definitions;
sultana
parents: 48829
diff changeset
   157
     S timed_test (interpretation_test (get_timeout @{context})) @{context} (TPTP_Syntax.get_file_list tptp_probs_dir))
47518
b2f209258621 more cleaning of tptp tests;
sultana
parents: 47516
diff changeset
   158
  else ()
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   159
\<close>
47512
b381d428a725 reorganised tptp testing thys
sultana
parents:
diff changeset
   160
62390
842917225d56 more canonical names
nipkow
parents: 58412
diff changeset
   161
end