src/HOL/TPTP/TPTP_Proof_Reconstruction_Test.thy
author wenzelm
Thu, 26 May 2016 17:51:22 +0200
changeset 63167 0909deb8059b
parent 62519 a564458f94db
child 65999 ee4cf96a9406
permissions -rw-r--r--
isabelle update_cartouches -c -t;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
59720
f893472fff31 proper headers;
wenzelm
parents: 58412
diff changeset
     1
(*  Title:      HOL/TPTP/TPTP_Proof_Reconstruction_Test.thy
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
     2
    Author:     Nik Sultana, Cambridge University Computer Laboratory
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
     3
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
     4
Various tests for the proof reconstruction module.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
     5
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
     6
NOTE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
     7
  - looks for THF proofs in the path indicated by $THF_PROOFS
55597
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
     8
  - these proofs are generated using LEO-II with the following
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
     9
    configuration choices: -po 1 -nux -nuc --expand_extuni
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    10
    You can simply filter LEO-II's output using the filter_proof
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    11
    script which is distributed with LEO-II.
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    12
*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    13
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    14
theory TPTP_Proof_Reconstruction_Test
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    15
imports TPTP_Test TPTP_Proof_Reconstruction
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    16
begin
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    17
55597
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    18
section "Main function"
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    19
text "This function wraps all the reconstruction-related functions. Simply
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    20
 point it at a THF proof produced by LEO-II (using the configuration described
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    21
 above), and it will try to reconstruct it into an Isabelle/HOL theorem.
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    22
 It also returns the theory (into which the axioms and definitions used in the
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    23
 proof have been imported), but note that you can get the theory value from
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    24
 the theorem value."
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    25
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
    26
ML \<open>
55597
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    27
  reconstruct_leo2 (Path.explode "$THF_PROOFS/NUM667^1.p.out") @{theory}
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
    28
\<close>
55597
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    29
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    30
section "Prep for testing the component functions"
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    31
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    32
declare [[
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    33
  tptp_trace_reconstruction = false,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    34
  tptp_test_all = false,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    35
  (* tptp_test_all = true, *)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    36
  tptp_test_timeout = 30,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    37
  (* tptp_max_term_size = 200 *)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    38
  tptp_max_term_size = 0
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    39
]]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    40
60926
0ccb5fb83c24 more standard options;
wenzelm
parents: 60925
diff changeset
    41
declare [[ML_exception_trace, ML_print_depth = 200]]
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    42
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    43
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    44
section "Importing proofs"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    45
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
    46
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    47
val probs =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    48
  (* "$THF_PROOFS/SYN991^1.p.out" *) (*lacks conjecture*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    49
  (* "$THF_PROOFS/SYO040^2.p.out" *)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    50
  (* "$THF_PROOFS/NUM640^1.p.out" *)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    51
  (* "$THF_PROOFS/SEU553^2.p.out" *)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    52
  (* "$THF_PROOFS/NUM665^1.p.out" *)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    53
  (* "$THF_PROOFS/SEV161^5.p.out" *)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    54
  (* "$THF_PROOFS/SET014^4.p.out" *)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    55
  "$THF_PROOFS/NUM667^1.p.out"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    56
  |> Path.explode
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    57
  |> single
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    58
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    59
val prob_names =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    60
  probs
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    61
  |> map (Path.base #> Path.implode #> TPTP_Problem_Name.Nonstandard)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
    62
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    63
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
    64
setup \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    65
  if test_all @{context} then I
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    66
  else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    67
    fold
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    68
     (fn path =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    69
       TPTP_Reconstruct.import_thm true [Path.dir path, Path.explode "$THF_PROOFS"] path leo2_on_load)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    70
     probs
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
    71
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    72
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    73
text "Display nicely."
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
    74
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    75
fun display_nicely ctxt (fms : TPTP_Reconstruct.formula_meaning list) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    76
  List.app (fn ((n, data) : TPTP_Reconstruct.formula_meaning) =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    77
    Pretty.writeln
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    78
      (Pretty.block
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    79
        [Pretty.str (n ^ " "),
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    80
         Syntax.pretty_term ctxt (#fmla data),
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    81
         Pretty.str (
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    82
          if is_none (#source_inf_opt data) then ""
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    83
          else ("\n\tannotation: " ^
60925
90659d0215bd prefer official @{make_string};
wenzelm
parents: 60649
diff changeset
    84
           @{make_string} (the (#source_inf_opt data : TPTP_Proof.source_info option))))])
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    85
    ) (rev fms);
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    86
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    87
(*FIXME hack for testing*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    88
fun test_fmla thy =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    89
  TPTP_Reconstruct.get_fmlas_of_prob thy (hd prob_names);
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    90
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    91
fun test_pannot thy =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    92
  TPTP_Reconstruct.get_pannot_of_prob thy (hd prob_names);
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    93
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    94
if test_all @{context} orelse prob_names = [] then ()
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    95
else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    96
  display_nicely @{context}
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    97
  (#meta (test_pannot @{theory}))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    98
(* To look at the original proof (i.e., before the proof transformations applied
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    99
   when the proof is loaded) replace previous line with:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   100
   (test_fmla @{theory}
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   101
    |> map TPTP_Reconstruct.structure_fmla_meaning)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   102
*)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   103
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   104
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   105
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   106
fun step_range_tester f_x f_exn ctxt prob_name from until =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   107
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   108
    val max =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   109
      case until of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   110
          SOME x => x
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   111
        | NONE =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   112
            if is_some Int.maxInt then the Int.maxInt else 999999
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   113
    fun test_step x =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   114
      if x > max then ()
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   115
      else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   116
        (f_x x;
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   117
         (interpret_leo2_inference ctxt prob_name (Int.toString x); ())
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   118
         handle e => f_exn e; (*FIXME naive. should let Interrupt through*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   119
         (*assumes that inferences are numbered consecutively*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   120
         test_step (x + 1))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   121
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   122
    test_step from
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   123
  end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   124
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   125
val step_range_tester_tracing =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   126
  step_range_tester
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   127
   (fn x => tracing ("@step " ^ Int.toString x))
60925
90659d0215bd prefer official @{make_string};
wenzelm
parents: 60649
diff changeset
   128
   (fn e => tracing ("!!" ^ @{make_string} e))
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   129
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   130
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   131
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   132
(*try to reconstruct each inference step*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   133
if test_all @{context} orelse prob_names = []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   134
orelse true (*NOTE currently disabled*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   135
then ()
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   136
else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   137
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   138
    (*FIXME not guaranteed to be the right nodes*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   139
    val heur_start = 3
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   140
    val heur_end =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   141
      hd (#meta (test_pannot @{theory}))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   142
      |> #1
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   143
      |> Int.fromString
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   144
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   145
    step_range_tester_tracing @{context} (hd prob_names) heur_start heur_end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   146
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   147
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   148
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   149
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   150
section "Building metadata and tactics"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   151
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   152
subsection "Building the skeleton"
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   153
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   154
if test_all @{context} orelse prob_names = [] then []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   155
else TPTP_Reconstruct.make_skeleton @{context} (test_pannot @{theory});
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   156
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   157
length it
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   158
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   159
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   160
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   161
subsection "The 'one shot' tactic approach"
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   162
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   163
val the_tactic =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   164
  if test_all @{context} then []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   165
  else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   166
    map (fn prob_name =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   167
      (TPTP_Reconstruct.naive_reconstruct_tac @{context} interpret_leo2_inference (* auto_based_reconstruction_tac *) (* oracle_based_reconstruction_tac *) prob_name))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   168
     prob_names;
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   169
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   170
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   171
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   172
subsection "The 'piecemeal' approach"
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   173
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   174
val the_tactics =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   175
  if test_all @{context} then []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   176
  else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   177
    map (fn prob_name =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   178
      TPTP_Reconstruct.naive_reconstruct_tacs interpret_leo2_inference (* auto_based_reconstruction_tac *) (* oracle_based_reconstruction_tac *) prob_name @{context})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   179
     prob_names;
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   180
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   181
56281
03c3d1a7c3b8 proper configuration option "ML_print_depth";
wenzelm
parents: 56265
diff changeset
   182
declare [[ML_print_depth = 2000]]
03c3d1a7c3b8 proper configuration option "ML_print_depth";
wenzelm
parents: 56265
diff changeset
   183
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   184
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   185
the_tactics
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   186
|> map (filter (fn (_, _, x) => is_none x)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   187
        #> map (fn (x, SOME y, _) => (x, cterm_of @{theory} y)))
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   188
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   189
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   190
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   191
section "Using metadata and tactics"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   192
text "There are various ways of testing the two ways (whole tactics or lists of tactics) of representing 'reconstructors'."
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   193
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   194
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   195
subsection "The 'one shot' tactic approach"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   196
text "First we test whole tactics."
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   197
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   198
(*produce thm*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   199
if test_all @{context} then []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   200
else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   201
  map (
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   202
    (* try *) (TPTP_Reconstruct.reconstruct @{context}
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   203
     (fn prob_name =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   204
       TPTP_Reconstruct.naive_reconstruct_tac @{context} interpret_leo2_inference prob_name
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   205
     (* oracle_based_reconstruction_tac *))))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   206
   prob_names
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   207
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   208
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   209
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   210
subsection "The 'piecemeal' approach"
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   211
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   212
fun attac n = List.nth (List.nth (the_tactics, 0), n) |> #3 |> the |> snd
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   213
fun attac_to n 0 = attac n
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   214
  | attac_to n m = attac n THEN attac_to (n + 1) (m - 1)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   215
fun shotac n = List.nth (List.nth (the_tactics, 0), n) |> #3 |> the |> fst
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   216
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   217
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   218
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   219
(*Given a list of reconstructed inferences (as in "the_tactics" above,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   220
  count the number of failures and successes, and list the failed inference
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   221
  reconstructions.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   222
fun evaluate_the_tactics [] acc = acc
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   223
  | evaluate_the_tactics ((node_no, (inf_name, inf_fmla, NONE)) :: xs) ((fai, suc), inf_list) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   224
      let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   225
        val score = (fai + 1, suc)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   226
        val index_info = get_index (fn (x, _) => if x = node_no then SOME true else NONE) inf_list
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   227
        val inf_list' =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   228
          case index_info of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   229
              NONE => (node_no, (inf_name, inf_fmla, 1)) :: inf_list
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   230
            | SOME (idx, _) =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   231
                nth_map idx (fn (node_no, (inf_name, inf_fmla, count)) => (node_no, (inf_name, inf_fmla, count + 1))) inf_list
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   232
      in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   233
        evaluate_the_tactics xs (score, inf_list')
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   234
      end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   235
  | evaluate_the_tactics ((_, (_, _, SOME _)) :: xs) ((fai, suc), inf_list) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   236
      evaluate_the_tactics xs ((fai, suc + 1), inf_list)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   237
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   238
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   239
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   240
text "Now we build a tactic by combining lists of tactics"
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   241
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   242
(*given a list of tactics to be applied in sequence (i.e., they
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   243
  follow a skeleton), we build a single tactic, interleaving
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   244
  some tracing info to help with debugging.*)
57772
7d9134b032b2 updated application of print_tac to take context parameter;
sultana
parents: 56467
diff changeset
   245
fun step_by_step_tacs ctxt verbose (thm_tacs : (thm * tactic) list) : tactic =
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   246
    let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   247
      fun interleave_tacs [] [] = all_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   248
        | interleave_tacs (tac1 :: tacs1) (tac2 :: tacs2) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   249
            EVERY [tac1, tac2]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   250
            THEN interleave_tacs tacs1 tacs2
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   251
      val thms_to_traceprint =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   252
        map (fn thm => fn st =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   253
              (*FIXME uses makestring*)
60925
90659d0215bd prefer official @{make_string};
wenzelm
parents: 60649
diff changeset
   254
              print_tac ctxt (@{make_string} thm) st)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   255
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   256
    in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   257
      if verbose then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   258
        ListPair.unzip thm_tacs
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   259
        |> apfst (fn thms => enumerate 1 thms)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   260
        |> apfst thms_to_traceprint
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   261
        |> uncurry interleave_tacs
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   262
      else EVERY (map #2 thm_tacs)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   263
    end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   264
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   265
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   266
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   267
(*apply step_by_step_tacs to all problems under test*)
57772
7d9134b032b2 updated application of print_tac to take context parameter;
sultana
parents: 56467
diff changeset
   268
fun narrated_tactics ctxt =
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   269
 map (map (#3 #> the)
57772
7d9134b032b2 updated application of print_tac to take context parameter;
sultana
parents: 56467
diff changeset
   270
      #> step_by_step_tacs ctxt false)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   271
   the_tactics;
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   272
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   273
(*produce thm*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   274
(*use narrated_tactics to reconstruct all problems under test*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   275
if test_all @{context} then []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   276
else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   277
  map (fn (prob_name, tac) =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   278
         TPTP_Reconstruct.reconstruct @{context}
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   279
           (fn _ => tac) prob_name)
57772
7d9134b032b2 updated application of print_tac to take context parameter;
sultana
parents: 56467
diff changeset
   280
    (ListPair.zip (prob_names, narrated_tactics @{context}))
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   281
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   282
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   283
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   284
subsection "Manually using 'piecemeal' approach"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   285
text "Another testing possibility involves manually creating a lemma
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   286
and running through the list of tactics generating to prove that lemma. The following code shows the goal of each problem under test, and then for each problem returns the list of tactics which can be invoked individually as shown below."
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   287
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   288
fun show_goal ctxt prob_name =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   289
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   290
    val thy = Proof_Context.theory_of ctxt
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   291
    val pannot = TPTP_Reconstruct.get_pannot_of_prob thy prob_name
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   292
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   293
    #meta pannot
58412
f65f11f4854c more standard Isabelle/ML operations;
wenzelm
parents: 58411
diff changeset
   294
    |> filter (fn (_, info) =>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   295
        #role info = TPTP_Syntax.Role_Conjecture)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   296
    |> hd
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   297
    |> snd |> #fmla
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   298
    |> cterm_of thy
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   299
  end;
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   300
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   301
if test_all @{context} then []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   302
else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   303
  map (show_goal @{context}) prob_names;
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   304
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   305
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   306
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   307
(*project out the list of tactics from "the_tactics"*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   308
val just_the_tacs  =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   309
 map (map (#3 #> the #> #2))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   310
   the_tactics;
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   311
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   312
map length just_the_tacs
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   313
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   314
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   315
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   316
(*like just_the_tacs, but extract the thms, to inspect their thys*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   317
val just_the_thms  =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   318
 map (map (#3 #> the #> #1))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   319
   the_tactics;
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   320
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   321
map length just_the_thms;
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   322
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   323
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   324
ML \<open>
60649
e007aa6a8aa2 more explicit use of context and elimination of Thm.theory_of_thm, although unclear (and untested?) situations remain;
wenzelm
parents: 59720
diff changeset
   325
val axms_of_thy =
e007aa6a8aa2 more explicit use of context and elimination of Thm.theory_of_thm, although unclear (and untested?) situations remain;
wenzelm
parents: 59720
diff changeset
   326
  `Theory.axioms_of
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   327
  #> apsnd cterm_of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   328
  #> swap
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   329
  #> apsnd (map snd)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   330
  #> uncurry map
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   331
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   332
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   333
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   334
(*Show the skeleton-level inference which is done by each element of just_the_tacs. This is useful when debugging using the technique shown next*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   335
if test_all @{context} orelse prob_names = [] then ()
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   336
else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   337
  the_tactics
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   338
  |> hd
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   339
  |> map #1
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   340
  |> TPTP_Reconstruct_Library.enumerate 0
60925
90659d0215bd prefer official @{make_string};
wenzelm
parents: 60649
diff changeset
   341
  |> List.app (@{make_string} #> writeln)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   342
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   343
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   344
ML \<open>
60649
e007aa6a8aa2 more explicit use of context and elimination of Thm.theory_of_thm, although unclear (and untested?) situations remain;
wenzelm
parents: 59720
diff changeset
   345
fun leo2_tac_wrap ctxt prob_name step i st =
e007aa6a8aa2 more explicit use of context and elimination of Thm.theory_of_thm, although unclear (and untested?) situations remain;
wenzelm
parents: 59720
diff changeset
   346
  rtac (interpret_leo2_inference ctxt prob_name step) i st
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   347
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   348
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   349
(*FIXME move these examples elsewhere*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   350
(*
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60926
diff changeset
   351
lemma "\<forall>(Xj::TPTP_Interpret.ind) Xk::TPTP_Interpret.ind.
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   352
        bnd_cCKB6_BLACK Xj Xk \<longrightarrow>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   353
        bnd_cCKB6_BLACK (bnd_s (bnd_s (bnd_s Xj))) (bnd_s Xk)"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   354
apply (tactic {*nth (nth just_the_tacs 0) 0*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   355
apply (tactic {*nth (nth just_the_tacs 0) 1*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   356
apply (tactic {*nth (nth just_the_tacs 0) 2*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   357
apply (tactic {*nth (nth just_the_tacs 0) 3*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   358
apply (tactic {*nth (nth just_the_tacs 0) 4*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   359
apply (tactic {*nth (nth just_the_tacs 0) 5*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   360
ML_prf "nth (hd the_tactics) 6"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   361
apply (tactic {*nth (nth just_the_tacs 0) 6*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   362
apply (tactic {*nth (nth just_the_tacs 0) 7*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   363
apply (tactic {*nth (nth just_the_tacs 0) 8*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   364
apply (tactic {*nth (nth just_the_tacs 0) 9*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   365
apply (tactic {*nth (nth just_the_tacs 0) 10*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   366
apply (tactic {*nth (nth just_the_tacs 0) 11*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   367
apply (tactic {*nth (nth just_the_tacs 0) 12*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   368
apply (tactic {*nth (nth just_the_tacs 0) 13*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   369
apply (tactic {*nth (nth just_the_tacs 0) 14*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   370
apply (tactic {*nth (nth just_the_tacs 0) 15*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   371
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   372
apply (tactic {*nth (nth just_the_tacs 0) 16*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   373
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   374
(*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   375
apply (tactic {*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   376
rtac (interpret_leo2_inference @{context} (hd prob_names) "8") 1
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   377
*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   378
apply (tactic {*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   379
rtac (interpret_leo2_inference @{context} (hd prob_names) "7") 1
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   380
*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   381
apply (tactic {*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   382
rtac (interpret_leo2_inference @{context} (hd prob_names) "6") 1
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   383
*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   384
(*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   385
apply (tactic {*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   386
rtac (interpret_leo2_inference @{context} (hd prob_names) "4") 1
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   387
*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   388
*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   389
*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   390
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   391
apply (tactic {*nth (nth just_the_tacs 0) 17*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   392
apply (tactic {*nth (nth just_the_tacs 0) 18*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   393
apply (tactic {*nth (nth just_the_tacs 0) 19*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   394
apply (tactic {*nth (nth just_the_tacs 0) 20*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   395
apply (tactic {*nth (nth just_the_tacs 0) 21*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   396
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   397
ML_prf "nth (hd the_tactics) 21"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   398
ML_prf "nth (hd the_tactics) 22"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   399
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   400
apply (tactic {*nth (nth just_the_tacs 0) 22*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   401
apply (tactic {*nth (nth just_the_tacs 0) 23*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   402
apply (tactic {*nth (nth just_the_tacs 0) 24*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   403
apply (tactic {*nth (nth just_the_tacs 0) 25*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   404
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   405
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   406
ML_prf "nth (hd the_tactics) 19"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   407
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   408
apply (tactic {*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   409
interpret_leo2_inference_wrap (hd prob_names) "8" 1
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   410
*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   411
apply (tactic {*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   412
interpret_leo2_inference_wrap (hd prob_names) "7" 1
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   413
*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   414
apply (tactic {*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   415
interpret_leo2_inference_wrap (hd prob_names) "6" 1
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   416
*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   417
apply (tactic {*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   418
interpret_leo2_inference_wrap (hd prob_names) "4" 1
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   419
*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   420
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   421
ML_prf "nth (hd the_tactics) 20"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   422
ML_prf "nth (hd the_tactics) 21"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   423
ML_prf "nth (hd the_tactics) 22"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   424
*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   425
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   426
(*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   427
lemma "bnd_powersetE1 \<longrightarrow>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   428
     bnd_sepInPowerset \<longrightarrow>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   429
     (\<forall>A Xphi. bnd_subset (bnd_dsetconstr A Xphi) A)"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   430
apply (tactic {*nth (nth just_the_tacs 0) 0*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   431
apply (tactic {*nth (nth just_the_tacs 0) 1*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   432
apply (tactic {*nth (nth just_the_tacs 0) 2*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   433
apply (tactic {*nth (nth just_the_tacs 0) 3*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   434
apply (tactic {*nth (nth just_the_tacs 0) 4*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   435
apply (tactic {*nth (nth just_the_tacs 0) 5*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   436
ML_prf "nth (hd the_tactics) 6"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   437
apply (tactic {*nth (nth just_the_tacs 0) 6*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   438
apply (tactic {*nth (nth just_the_tacs 0) 7*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   439
apply (tactic {*nth (nth just_the_tacs 0) 8*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   440
apply (tactic {*nth (nth just_the_tacs 0) 9*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   441
apply (tactic {*nth (nth just_the_tacs 0) 10*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   442
apply (tactic {*nth (nth just_the_tacs 0) 11*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   443
apply (tactic {*nth (nth just_the_tacs 0) 12*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   444
apply (tactic {*nth (nth just_the_tacs 0) 13*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   445
apply (tactic {*nth (nth just_the_tacs 0) 14*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   446
apply (tactic {*nth (nth just_the_tacs 0) 15*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   447
apply (tactic {*nth (nth just_the_tacs 0) 16*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   448
apply (tactic {*nth (nth just_the_tacs 0) 17*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   449
apply (tactic {*nth (nth just_the_tacs 0) 18*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   450
apply (tactic {*nth (nth just_the_tacs 0) 19*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   451
apply (tactic {*nth (nth just_the_tacs 0) 20*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   452
apply (tactic {*nth (nth just_the_tacs 0) 21*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   453
apply (tactic {*nth (nth just_the_tacs 0) 22*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   454
apply (tactic {*nth (nth just_the_tacs 0) 23*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   455
apply (tactic {*nth (nth just_the_tacs 0) 24*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   456
apply (tactic {*nth (nth just_the_tacs 0) 25*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   457
(* apply (tactic {*nth (nth just_the_tacs 0) 26*}) *)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   458
ML_prf "nth (hd the_tactics) 26"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   459
apply (subgoal_tac "(\<not> (\<forall>A Xphi. bnd_subset (bnd_dsetconstr A Xphi) A)) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   460
       True \<Longrightarrow>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   461
       (\<not> bnd_subset (bnd_dsetconstr bnd_sK1 bnd_sK2) bnd_sK1) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   462
       True")
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   463
prefer 2
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   464
apply (thin_tac "(bnd_powersetE1 \<longrightarrow>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   465
      bnd_sepInPowerset \<longrightarrow>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   466
      (\<forall>A Xphi. bnd_subset (bnd_dsetconstr A Xphi) A)) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   467
     False")
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   468
apply (tactic {*extcnf_combined_simulator_tac (hd prob_names) 1*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   469
apply (tactic {*extcnf_combined_simulator_tac (hd prob_names) 1*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   470
apply (tactic {*extcnf_combined_simulator_tac (hd prob_names) 1*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   471
apply (tactic {*extcnf_combined_simulator_tac (hd prob_names) 1*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   472
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   473
apply simp
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   474
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   475
(* apply (tactic {*nth (nth just_the_tacs 0) 26*}) *)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   476
apply (tactic {*nth (nth just_the_tacs 0) 27*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   477
apply (tactic {*nth (nth just_the_tacs 0) 28*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   478
apply (tactic {*nth (nth just_the_tacs 0) 29*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   479
apply (tactic {*nth (nth just_the_tacs 0) 30*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   480
apply (tactic {*nth (nth just_the_tacs 0) 31*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   481
apply (tactic {*nth (nth just_the_tacs 0) 32*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   482
apply (tactic {*nth (nth just_the_tacs 0) 33*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   483
apply (tactic {*nth (nth just_the_tacs 0) 34*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   484
apply (tactic {*nth (nth just_the_tacs 0) 35*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   485
apply (tactic {*nth (nth just_the_tacs 0) 36*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   486
apply (tactic {*nth (nth just_the_tacs 0) 37*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   487
apply (tactic {*nth (nth just_the_tacs 0) 38*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   488
apply (tactic {*nth (nth just_the_tacs 0) 39*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   489
apply (tactic {*nth (nth just_the_tacs 0) 40*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   490
apply (tactic {*nth (nth just_the_tacs 0) 41*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   491
apply (tactic {*nth (nth just_the_tacs 0) 42*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   492
apply (tactic {*nth (nth just_the_tacs 0) 43*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   493
apply (tactic {*nth (nth just_the_tacs 0) 44*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   494
apply (tactic {*nth (nth just_the_tacs 0) 45*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   495
apply (tactic {*nth (nth just_the_tacs 0) 46*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   496
apply (tactic {*nth (nth just_the_tacs 0) 47*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   497
apply (tactic {*nth (nth just_the_tacs 0) 48*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   498
apply (tactic {*nth (nth just_the_tacs 0) 49*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   499
apply (tactic {*nth (nth just_the_tacs 0) 50*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   500
apply (tactic {*nth (nth just_the_tacs 0) 51*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   501
done
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   502
*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   503
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   504
(*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   505
We can use just_the_tacs as follows:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   506
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   507
(this is from SEV012^5.p.out)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   508
lemma "((\<forall>(Xx :: bool) (Xy :: bool). True \<longrightarrow> True) \<and>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   509
      (\<forall>(Xx :: bool) (Xy :: bool) (Xz :: bool). True \<and> True \<longrightarrow> True)) \<and>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   510
     (\<lambda>(Xx :: bool) (Xy :: bool). True) = (\<lambda>Xx Xy. True)"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   511
apply (tactic {*nth (nth just_the_tacs 0) 0*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   512
apply (tactic {*nth (nth just_the_tacs 0) 1*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   513
apply (tactic {*nth (nth just_the_tacs 0) 2*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   514
apply (tactic {*nth (nth just_the_tacs 0) 3*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   515
apply (tactic {*nth (nth just_the_tacs 0) 4*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   516
apply (tactic {*nth (nth just_the_tacs 0) 5*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   517
ML_prf "nth (hd the_tactics) 6"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   518
apply (tactic {*nth (nth just_the_tacs 0) 6*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   519
apply (tactic {*nth (nth just_the_tacs 0) 7*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   520
apply (tactic {*nth (nth just_the_tacs 0) 8*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   521
apply (tactic {*nth (nth just_the_tacs 0) 9*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   522
apply (tactic {*nth (nth just_the_tacs 0) 10*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   523
apply (tactic {*nth (nth just_the_tacs 0) 11*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   524
apply (tactic {*nth (nth just_the_tacs 0) 12*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   525
apply (tactic {*nth (nth just_the_tacs 0) 13*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   526
apply (tactic {*nth (nth just_the_tacs 0) 14*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   527
apply (tactic {*nth (nth just_the_tacs 0) 15*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   528
apply (tactic {*nth (nth just_the_tacs 0) 16*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   529
apply (tactic {*nth (nth just_the_tacs 0) 17*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   530
apply (tactic {*nth (nth just_the_tacs 0) 18*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   531
apply (tactic {*nth (nth just_the_tacs 0) 19*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   532
apply (tactic {*nth (nth just_the_tacs 0) 20*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   533
apply (tactic {*nth (nth just_the_tacs 0) 21*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   534
apply (tactic {*nth (nth just_the_tacs 0) 22*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   535
done
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   536
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   537
(*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   538
We could also use previous definitions directly,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   539
e.g. the following should prove the goal at a go:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   540
- apply (tactic {*narrated_tactics |> hd |> hd*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   541
- apply (tactic {*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   542
    TPTP_Reconstruct.naive_reconstruct_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   543
     interpret_leo2_inference
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   544
     (hd prob_names)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   545
     @{context}*})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   546
(Note that the previous two methods don't work in this
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   547
 "lemma" testing mode, not sure why. The previous methods
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   548
 (producing the thm values directly) should work though.)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   549
*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   550
*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   551
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   552
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   553
section "Testing against benchmark"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   554
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   555
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   556
(*if reconstruction_info value is NONE then a big error must have occurred*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   557
type reconstruction_info =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   558
  ((int(*no of failures*) * int(*no of successes*)) *
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   559
  (TPTP_Reconstruct.rolling_stock * term option(*inference formula*) * int (*number of times the inference occurs in the skeleton*)) list) option
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   560
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   561
datatype proof_contents =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   562
    No_info
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   563
  | Empty
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   564
  | Nonempty of reconstruction_info
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   565
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   566
(*To make output less cluttered in whole-run tests*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   567
fun erase_inference_fmlas (Nonempty (SOME (outline, inf_info))) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   568
      Nonempty (SOME (outline, map (fn (inf_name, _, count) => (inf_name, NONE, count)) inf_info))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   569
  | erase_inference_fmlas x = x
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   570
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   571
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   572
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   573
(*Report on how many inferences in a proof are reconstructed, and give some
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   574
  info about the inferences for which reconstruction failed.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   575
fun test_partial_reconstruction thy prob_file =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   576
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   577
    val prob_name =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   578
      (Path.base #> Path.implode #> TPTP_Problem_Name.Nonstandard) prob_file
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   579
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   580
    val thy' =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   581
      try
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   582
       (TPTP_Reconstruct.import_thm
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   583
        true
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   584
        [Path.dir prob_file, Path.explode "$TPTP"]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   585
        prob_file leo2_on_load)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   586
       thy
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   587
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   588
    val ctxt' =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   589
      if is_some thy' then SOME (Proof_Context.init_global (the thy')) else NONE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   590
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   591
    (*to test if proof is empty*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   592
    val fms =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   593
      if is_some thy'
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   594
      then SOME (TPTP_Reconstruct.get_fmlas_of_prob (the thy') prob_name)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   595
      else NONE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   597
    val the_tactics =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   598
      if is_some thy' then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   599
        SOME (TPTP_Reconstruct.naive_reconstruct_tacs (* metis_based_reconstruction_tac *)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   600
interpret_leo2_inference (* auto_based_reconstruction_tac *) (* oracle_based_reconstruction_tac *) prob_name (the ctxt'))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   601
      else NONE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   602
60925
90659d0215bd prefer official @{make_string};
wenzelm
parents: 60649
diff changeset
   603
(* val _ = tracing ("tt=" ^ @{make_string} the_tactics) *)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   604
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   605
    val skeleton =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   606
      if is_some thy' then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   607
        SOME (TPTP_Reconstruct.make_skeleton (the ctxt')
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   608
              (TPTP_Reconstruct.get_pannot_of_prob (the thy') prob_name))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   609
      else NONE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   610
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   611
    val skeleton_and_tactics =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   612
      if is_some thy' then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   613
        SOME (ListPair.zip (the skeleton, the the_tactics))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   614
      else NONE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   615
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   616
    val result =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   617
      if is_some thy' then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   618
        SOME (evaluate_the_tactics (the skeleton_and_tactics)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   619
              ((0, 0), []))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   620
      else NONE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   621
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   622
    (*strip node names*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   623
    val result' =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   624
      if is_some result then SOME (apsnd (map #2) (the result)) else NONE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   625
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   626
    if is_some fms andalso List.null (the fms) then Empty
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   627
    else Nonempty result'
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   628
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   629
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   630
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   631
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   632
  (*default timeout is 1 min*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   633
  fun reconstruct timeout light_output file thy =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   634
    let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   635
      val timer = Timer.startRealTimer ()
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   636
    in
62519
a564458f94db tuned signature -- clarified modules;
wenzelm
parents: 62390
diff changeset
   637
      Timeout.apply (Time.fromSeconds (if timeout = 0 then 60 else timeout))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   638
       (test_partial_reconstruction thy
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   639
        #> light_output ? erase_inference_fmlas
60925
90659d0215bd prefer official @{make_string};
wenzelm
parents: 60649
diff changeset
   640
        #> @{make_string} (* FIXME *)
90659d0215bd prefer official @{make_string};
wenzelm
parents: 60649
diff changeset
   641
        #> (fn s => report (Proof_Context.init_global thy) (@{make_string} file ^ " === " ^ s ^
90659d0215bd prefer official @{make_string};
wenzelm
parents: 60649
diff changeset
   642
             " t=" ^ (Timer.checkRealTimer timer |> Time.toMilliseconds |> @{make_string}))))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   643
       file
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   644
    end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   645
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   646
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   647
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   648
  (*this version of "reconstruct" builds theorems, instead of lists of reconstructed inferences*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   649
  (*default timeout is 1 min*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   650
  fun reconstruct timeout file thy =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   651
    let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   652
      val timer = Timer.startRealTimer ()
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   653
      val thy' =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   654
        TPTP_Reconstruct.import_thm true
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   655
         [Path.dir file, Path.explode "$TPTP"]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   656
         file leo2_on_load thy
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   657
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   658
      val ctxt = Proof_Context.init_global thy' (*FIXME pass ctxt instead of thy*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   659
      val prob_name =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   660
        file
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   661
        |> Path.base
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   662
        |> Path.implode
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   663
        |> TPTP_Problem_Name.Nonstandard
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   664
    in
62519
a564458f94db tuned signature -- clarified modules;
wenzelm
parents: 62390
diff changeset
   665
      Timeout.apply (Time.fromSeconds (if timeout = 0 then 60 else timeout))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   666
       (fn prob_name =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   667
        (can
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   668
          (TPTP_Reconstruct.reconstruct ctxt (fn prob_name =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   669
            TPTP_Reconstruct.naive_reconstruct_tac ctxt interpret_leo2_inference prob_name (* oracle_based_reconstruction_tac *))) prob_name )
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   670
       |> (fn s => report ctxt (Path.print file ^ " === " ^ Bool.toString s ^
60925
90659d0215bd prefer official @{make_string};
wenzelm
parents: 60649
diff changeset
   671
             " t=" ^ (Timer.checkRealTimer timer |> Time.toMilliseconds |> @{make_string}))))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   672
       prob_name
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   673
    end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   674
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   675
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   676
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   677
  fun reconstruction_test timeout ctxt =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   678
    test_fn ctxt
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   679
     (fn file => reconstruct timeout file (Proof_Context.theory_of ctxt))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   680
     "reconstructor"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   681
     ()
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   682
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   683
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   684
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   685
datatype examination_results =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   686
    Whole_proof of string(*filename*) * proof_contents
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   687
  | Specific_rule of string(*filename*) * string(*inference rule*) * term option list
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   688
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   689
(*Look out for failures reconstructing a particular inference rule*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   690
fun filter_failures inference_name (Whole_proof (filename, results)) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   691
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   692
    val filtered_results =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   693
      case results of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   694
          Nonempty (SOME results') =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   695
            #2 results'
58411
wenzelm
parents: 57772
diff changeset
   696
            |> maps (fn (stock as TPTP_Reconstruct.Annotated_step (_, inf_name), inf_fmla, _) =>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   697
                 if inf_name = inference_name then [inf_fmla] else [])
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   698
        | _ => []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   699
  in Specific_rule (filename, inference_name, filtered_results) end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   700
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   701
(*Returns detailed info about a proof-reconstruction attempt.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   702
  If rule_name is specified then the related failed inferences
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   703
  are returned, otherwise all failed inferences are returned.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   704
fun examine_failed_inferences ctxt filename rule_name =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   705
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   706
    val thy = Proof_Context.theory_of ctxt
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   707
    val prob_file = Path.explode filename
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   708
    val results =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   709
      if test_all ctxt then No_info
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   710
      else test_partial_reconstruction thy prob_file
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   711
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   712
    Whole_proof (filename, results)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   713
    |> is_some rule_name ? (fn x =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   714
                             filter_failures (the rule_name) x)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   715
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   716
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   717
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   718
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   719
exception NONSENSE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   720
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   721
fun annotation_or_id (TPTP_Reconstruct.Step n) = n
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   722
  | annotation_or_id (TPTP_Reconstruct.Annotated_step (n, anno)) = anno
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   723
  | annotation_or_id TPTP_Reconstruct.Assumed = "assumption"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   724
  | annotation_or_id TPTP_Reconstruct.Unconjoin = "conjI"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   725
  | annotation_or_id TPTP_Reconstruct.Caboose = "(end)"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   726
  | annotation_or_id (TPTP_Reconstruct.Synth_step s) = s
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   727
  | annotation_or_id (TPTP_Reconstruct.Split (split_node, soln_node, _)) = "split_at " ^ split_node ^ " " ^ soln_node;
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   728
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   729
fun count_failures (Whole_proof (_, No_info)) = raise NONSENSE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   730
  | count_failures (Whole_proof (_, Empty)) = raise NONSENSE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   731
  | count_failures (Whole_proof (_, Nonempty NONE)) = raise NONSENSE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   732
  | count_failures (Whole_proof (_, Nonempty (SOME (((n, _), _))))) = n
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   733
  | count_failures (Specific_rule (_, _, t)) = length t
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   734
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   735
fun pre_classify_failures [] alist = alist
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   736
  | pre_classify_failures ((stock, _, _) :: xs) alist =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   737
      let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   738
        val inf = annotation_or_id stock
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   739
        val count = AList.lookup (op =) alist inf
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   740
      in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   741
        if is_none count
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   742
        then pre_classify_failures xs ((inf, 1) :: alist)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   743
        else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   744
          pre_classify_failures xs
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   745
           (AList.update (op =) (inf, the count + 1) alist)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   746
      end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   747
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   748
fun classify_failures (Whole_proof (_, Nonempty (SOME (((_, _), inferences))))) = pre_classify_failures inferences []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   749
  | classify_failures (Specific_rule (_, rule, t)) = [(rule, length t)]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   750
  | classify_failures _ = raise NONSENSE
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   751
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   752
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   753
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   754
val regressions = map (fn s => "$THF_PROOFS/" ^ s)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   755
  ["SEV405^5.p.out",
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   756
   (*"SYO377^5.p.out", Always seems to raise Interrupt on my laptop -- probably because node 475 has lots of premises*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   757
   "PUZ031^5.p.out",
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   758
   "ALG001^5.p.out",
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   759
   "SYO238^5.p.out",
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   760
   (*"SEV158^5.p.out", This is big*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   761
   "SYO285^5.p.out",
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   762
   "../SYO285^5.p.out_reduced",
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   763
   (* "SYO225^5.p.out", This is big*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   764
   "SYO291^5.p.out",
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   765
   "SET669^3.p.out",
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   766
   "SEV233^5.p.out",
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   767
   (*"SEU511^1.p.out", This is big*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   768
   "SEV161^5.p.out",
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   769
   "SEV012^5.p.out",
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   770
   "SYO035^1.p.out",
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   771
   "SYO291^5.p.out",
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   772
   "SET741^4.p.out", (*involves both definitions and contorted splitting. has nice graph.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   773
   "SEU548^2.p.out",
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   774
   "SEU513^2.p.out",
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   775
   "SYO006^1.p.out",
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   776
   "SYO371^5.p.out" (*has contorted splitting, like SYO006^1.p.out, but doesn't involve definitions*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   777
  ]
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   778
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   779
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   780
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   781
val experiment = examine_failed_inferences @{context}
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   782
  (List.last regressions) NONE;
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   783
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   784
(*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   785
val experiment_focus =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   786
  filter_failures "extcnf_combined" experiment;
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   787
*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   788
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   789
(*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   790
count_failures experiment_focus
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   791
classify_failures experiment
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   792
*)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   793
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   794
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   795
text "Run reconstruction on all problems in a benchmark (provided via a script)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   796
and report on partial success."
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   797
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   798
declare [[
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   799
  tptp_test_all = true,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   800
  tptp_test_timeout = 10
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   801
]]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   802
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   803
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   804
  (*problem source*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   805
  val tptp_probs_dir =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   806
    Path.explode "$THF_PROOFS"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   807
    |> Path.expand;
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   808
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   809
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   810
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   811
  if test_all @{context} then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   812
    (report @{context} "Reconstructing proofs";
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   813
    S timed_test (reconstruction_test (get_timeout @{context})) @{context} (TPTP_Syntax.get_file_list tptp_probs_dir))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   814
  else ()
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62519
diff changeset
   815
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   816
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   817
(*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   818
Debugging strategy:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   819
  1) get list of all proofs
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   820
  2) order by size
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   821
  3) try to construct each in turn, given some timeout
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   822
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   823
Use this to find the smallest failure, then debug that.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   824
*)
62390
842917225d56 more canonical names
nipkow
parents: 61076
diff changeset
   825
end