src/HOL/TPTP/TPTP_Proof_Reconstruction.thy
author wenzelm
Mon, 20 May 2024 15:43:51 +0200
changeset 80182 29f2b8ff84f3
parent 78727 1b052426a2b7
child 80636 4041e7c8059d
permissions -rw-r--r--
proper support for "isabelle update -D DIR": avoid accidental exclusion of select_dirs (amending e5dafe9e120f);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
     1
(*  Title:      HOL/TPTP/TPTP_Proof_Reconstruction.thy
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
Proof reconstruction for Leo-II.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
     5
55597
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
     6
USAGE:
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
     7
* Simple call the "reconstruct_leo2" function.
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
     8
* For more advanced use, you could use the component functions used in
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
     9
  "reconstruct_leo2" -- see TPTP_Proof_Reconstruction_Test.thy for
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    10
  examples of this.
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    11
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    12
This file contains definitions describing how to interpret LEO-II's
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    13
calculus in Isabelle/HOL, as well as more general proof-handling
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    14
functions. The definitions in this file serve to build an intermediate
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    15
proof script which is then evaluated into a tactic -- both these steps
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    16
are independent of LEO-II, and are defined in the TPTP_Reconstruct SML
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    17
module.
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    18
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    19
CONFIG:
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    20
The following attributes are mainly useful for debugging:
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    21
  tptp_unexceptional_reconstruction |
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    22
  unexceptional_reconstruction      |-- when these are true, a low-level exception
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    23
                                        is allowed to float to the top (instead of
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    24
                                        triggering a higher-level exception, or
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    25
                                        simply indicating that the reconstruction failed).
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    26
  tptp_max_term_size                --- fail of a term exceeds this size. "0" is taken
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    27
                                        to mean infinity.
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    28
  tptp_informative_failure          |
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    29
  informative_failure               |-- produce more output during reconstruction.
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    30
  tptp_trace_reconstruction         |
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    31
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    32
There are also two attributes, independent of the code here, that
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    33
influence the success of reconstruction: blast_depth_limit and
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    34
unify_search_bound. These are documented in their respective modules,
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    35
but in summary, if unify_search_bound is increased then we can
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    36
handle larger terms (at the cost of performance), since the unification
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    37
engine takes longer to give up the search; blast_depth_limit is
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    38
a limit on proof search performed by Blast. Blast is used for
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    39
the limited proof search that needs to be done to interpret
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    40
instances of LEO-II's inference rules.
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
    41
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    42
TODO:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    43
  use RemoveRedundantQuantifications instead of the ad hoc use of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    44
   remove_redundant_quantification_in_lit and remove_redundant_quantification
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    45
*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    46
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    47
theory TPTP_Proof_Reconstruction
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    48
imports TPTP_Parser TPTP_Interpret
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    49
(* keywords "import_leo2_proof" :: thy_decl *) (*FIXME currently unused*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    50
begin
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    51
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    52
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    53
section "Setup"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    54
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
    55
ML \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
    56
  val tptp_unexceptional_reconstruction = Attrib.setup_config_bool \<^binding>\<open>tptp_unexceptional_reconstruction\<close> (K false)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    57
  fun unexceptional_reconstruction ctxt = Config.get ctxt tptp_unexceptional_reconstruction
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
    58
  val tptp_informative_failure = Attrib.setup_config_bool \<^binding>\<open>tptp_informative_failure\<close> (K false)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    59
  fun informative_failure ctxt = Config.get ctxt tptp_informative_failure
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
    60
  val tptp_trace_reconstruction = Attrib.setup_config_bool \<^binding>\<open>tptp_trace_reconstruction\<close> (K false)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
    61
  val tptp_max_term_size = Attrib.setup_config_int \<^binding>\<open>tptp_max_term_size\<close> (K 0) (*0=infinity*)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    62
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    63
  fun exceeds_tptp_max_term_size ctxt size =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    64
    let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    65
      val max = Config.get ctxt tptp_max_term_size
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    66
    in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    67
      if max = 0 then false
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    68
      else size > max
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    69
    end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
    70
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    71
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    72
(*FIXME move to TPTP_Proof_Reconstruction_Test_Units*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    73
declare [[
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    74
  tptp_unexceptional_reconstruction = false, (*NOTE should be "false" while testing*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    75
  tptp_informative_failure = true
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    76
]]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    77
78727
1b052426a2b7 avoid accidental 'handle' of interrupts;
wenzelm
parents: 77879
diff changeset
    78
ML \<open>
1b052426a2b7 avoid accidental 'handle' of interrupts;
wenzelm
parents: 77879
diff changeset
    79
exception UNSUPPORTED_ROLE
1b052426a2b7 avoid accidental 'handle' of interrupts;
wenzelm
parents: 77879
diff changeset
    80
exception INTERPRET_INFERENCE
1b052426a2b7 avoid accidental 'handle' of interrupts;
wenzelm
parents: 77879
diff changeset
    81
\<close>
1b052426a2b7 avoid accidental 'handle' of interrupts;
wenzelm
parents: 77879
diff changeset
    82
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69597
diff changeset
    83
ML_file \<open>TPTP_Parser/tptp_reconstruct_library.ML\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    84
ML "open TPTP_Reconstruct_Library"
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69597
diff changeset
    85
ML_file \<open>TPTP_Parser/tptp_reconstruct.ML\<close>
55596
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 fudge*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    88
declare [[
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    89
  blast_depth_limit = 10,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    90
  unify_search_bound = 5
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    91
]]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    92
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
section "Proof reconstruction"
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
    95
text \<open>There are two parts to proof reconstruction:
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    96
\begin{itemize}
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    97
  \item interpreting the inferences
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    98
  \item building the skeleton, which indicates how to compose
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
    99
    individual inferences into subproofs, and then compose the
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   100
    subproofs to give the proof).
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   101
\end{itemize}
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   102
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   103
One step detects unsound inferences, and the other step detects
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   104
unsound composition of inferences.  The two parts can be weakly
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   105
coupled. They rely on a "proof index" which maps nodes to the
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   106
inference information. This information consists of the (usually
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   107
prover-specific) name of the inference step, and the Isabelle
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   108
formalisation of the inference as a term. The inference interpretation
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   109
then maps these terms into meta-theorems, and the skeleton is used to
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   110
compose the inference-level steps into a proof.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   111
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   112
Leo2 operates on conjunctions of clauses. Each Leo2 inference has the
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   113
following form, where Cx are clauses:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   114
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   115
           C1 && ... && Cn
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   116
          -----------------
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   117
          C'1 && ... && C'n
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   118
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   119
Clauses consist of disjunctions of literals (shown as Px below), and might
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   120
have a prefix of !-bound variables, as shown below.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   121
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   122
  ! X... { P1 || ... || Pn}
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   123
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   124
Literals are usually assigned a polarity, but this isn't always the
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   125
case; you can come across inferences looking like this (where A is an
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   126
object-level formula):
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   127
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   128
             F
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   129
          --------
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   130
          F = true
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   131
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   132
The symbol "||" represents literal-level disjunction and "&&" is
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   133
clause-level conjunction. Rules will typically lift formula-level
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   134
conjunctions; for instance the following rule lifts object-level
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   135
disjunction:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   136
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   137
          {    (A | B) = true    || ... } && ...
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   138
          --------------------------------------
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   139
          { A = true || B = true || ... } && ...
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   140
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   141
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   142
Using this setup, efficiency might be gained by only interpreting
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   143
inferences once, merging identical inference steps, and merging
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   144
identical subproofs into single inferences thus avoiding some effort.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   145
We can also attempt to minimising proof search when interpreting
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   146
inferences.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   147
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   148
It is hoped that this setup can target other provers by modifying the
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   149
clause representation to fit them, and adapting the inference
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   150
interpretation to handle the rules used by the prover. It should also
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   151
facilitate composing together proofs found by different provers.
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   152
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   153
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   154
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   155
subsection "Instantiation"
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
lemma polar_allE [rule_format]:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   158
  "\<lbrakk>(\<forall>x. P x) = True; (P x) = True \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   159
  "\<lbrakk>(\<exists>x. P x) = False; (P x) = False \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   160
by auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   161
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   162
lemma polar_exE [rule_format]:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   163
  "\<lbrakk>(\<exists>x. P x) = True; \<And>x. (P x) = True \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   164
  "\<lbrakk>(\<forall>x. P x) = False; \<And>x. (P x) = False \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   165
by auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   166
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   167
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   168
(*This carries out an allE-like rule but on (polarised) literals.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   169
 Instead of yielding a free variable (which is a hell for the
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   170
 matcher) it seeks to use one of the subgoals' parameters.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   171
 This ought to be sufficient for emulating extcnf_combined,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   172
 but note that the complexity of the problem can be enormous.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   173
fun inst_parametermatch_tac ctxt thms i = fn st =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   174
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   175
    val gls =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59533
diff changeset
   176
      Thm.prop_of st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   177
      |> Logic.strip_horn
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   178
      |> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   179
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   180
    val parameters =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   181
      if null gls then []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   182
      else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   183
        rpair (i - 1) gls
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   184
        |> uncurry nth
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   185
        |> strip_top_all_vars []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   186
        |> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   187
        |> map fst (*just get the parameter names*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   188
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   189
    if null parameters then no_tac st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   190
    else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   191
      let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   192
        fun instantiate param =
59780
23b67731f4f0 support 'for' fixes in rule_tac etc.;
wenzelm
parents: 59763
diff changeset
   193
          (map (Rule_Insts.eres_inst_tac ctxt [((("x", 0), Position.none), param)] []) thms
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   194
                   |> FIRST')
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   195
        val attempts = map instantiate parameters
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   196
      in
67405
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67404
diff changeset
   197
        (fold (curry (op APPEND')) attempts (K no_tac)) i st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   198
      end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   199
  end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   200
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   201
(*Attempts to use the polar_allE theorems on a specific subgoal.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   202
fun forall_pos_tac ctxt = inst_parametermatch_tac ctxt @{thms polar_allE}
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   203
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   204
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   205
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   206
(*This is similar to inst_parametermatch_tac, but prefers to
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   207
  match variables having identical names. Logically, this is
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   208
  a hack. But it reduces the complexity of the problem.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   209
fun nominal_inst_parametermatch_tac ctxt thm i = fn st =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   210
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   211
    val gls =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59533
diff changeset
   212
      Thm.prop_of st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   213
      |> Logic.strip_horn
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   214
      |> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   215
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   216
    val parameters =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   217
      if null gls then []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   218
      else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   219
        rpair (i - 1) gls
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   220
        |> uncurry nth
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   221
        |> strip_top_all_vars []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   222
        |> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   223
        |> map fst (*just get the parameter names*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   224
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   225
    if null parameters then no_tac st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   226
    else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   227
      let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   228
        fun instantiates param =
59780
23b67731f4f0 support 'for' fixes in rule_tac etc.;
wenzelm
parents: 59763
diff changeset
   229
          Rule_Insts.eres_inst_tac ctxt [((("x", 0), Position.none), param)] [] thm
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   230
59533
e9ffe89a20a4 proper context;
wenzelm
parents: 59498
diff changeset
   231
        val quantified_var = head_quantified_variable ctxt i st
55596
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
        if is_none quantified_var then no_tac st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   234
        else
67405
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67404
diff changeset
   235
          if member (op =) parameters (the quantified_var |> fst) then
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   236
            instantiates (the quantified_var |> fst) i st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   237
          else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   238
            K no_tac i st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   239
      end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   240
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   241
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   242
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   243
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   244
subsection "Prefix massaging"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   245
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   246
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   247
exception NO_GOALS
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   248
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   249
(*Get quantifier prefix of the hypothesis and conclusion, reorder
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   250
  the hypothesis' quantifiers to have the ones appearing in the
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   251
  conclusion first.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   252
fun canonicalise_qtfr_order ctxt i = fn st =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   253
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   254
    val gls =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59533
diff changeset
   255
      Thm.prop_of st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   256
      |> Logic.strip_horn
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   257
      |> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   258
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   259
    if null gls then raise NO_GOALS
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   260
    else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   261
      let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   262
        val (params, (hyp_clause, conc_clause)) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   263
          rpair (i - 1) gls
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   264
          |> uncurry nth
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   265
          |> strip_top_all_vars []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   266
          |> apsnd Logic.dest_implies
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   267
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   268
        val (hyp_quants, hyp_body) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   269
          HOLogic.dest_Trueprop hyp_clause
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   270
          |> strip_top_All_vars
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   271
          |> apfst rev
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
        val conc_quants =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   274
          HOLogic.dest_Trueprop conc_clause
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   275
          |> strip_top_All_vars
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   276
          |> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   277
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   278
        val new_hyp =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   279
          (* fold absfree new_hyp_prefix hyp_body *)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   280
          (*HOLogic.list_all*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   281
          fold_rev (fn (v, ty) => fn t => HOLogic.mk_all (v, ty, t))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   282
           (prefix_intersection_list
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   283
             hyp_quants conc_quants)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   284
           hyp_body
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   285
          |> HOLogic.mk_Trueprop
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   286
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   287
         val thm = Goal.prove ctxt [] []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   288
           (Logic.mk_implies (hyp_clause, new_hyp))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   289
           (fn _ =>
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   290
              (REPEAT_DETERM (HEADGOAL (resolve_tac ctxt @{thms allI})))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   291
              THEN (REPEAT_DETERM
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   292
                    (HEADGOAL
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   293
                     (nominal_inst_parametermatch_tac ctxt @{thm allE})))
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   294
              THEN HEADGOAL (assume_tac ctxt))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   295
      in
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   296
        dresolve_tac ctxt [thm] i st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   297
      end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   298
    end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   299
\<close>
55596
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
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   302
subsection "Some general rules and congruences"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   303
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   304
(*this isn't an actual rule used in Leo2, but it seems to be
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   305
  applied implicitly during some Leo2 inferences.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   306
lemma polarise: "P ==> P = True" by auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   307
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   308
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   309
fun is_polarised t =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   310
  (TPTP_Reconstruct.remove_polarity true t; true)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   311
  handle TPTP_Reconstruct.UNPOLARISED _ => false
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   312
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   313
fun polarise_subgoal_hyps ctxt =
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   314
  COND' (SOME #> TERMPRED is_polarised (fn _ => true)) (K no_tac) (dresolve_tac ctxt @{thms polarise})
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   315
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   316
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   317
lemma simp_meta [rule_format]:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   318
  "(A --> B) == (~A | B)"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   319
  "(A | B) | C == A | B | C"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   320
  "(A & B) & C == A & B & C"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   321
  "(~ (~ A)) == A"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   322
  (* "(A & B) == (~ (~A | ~B))" *)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   323
  "~ (A & B) == (~A | ~B)"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   324
  "~(A | B) == (~A) & (~B)"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   325
by auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   326
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   327
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   328
subsection "Emulation of Leo2's inference rules"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   329
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   330
(*this is not included in simp_meta since it would make a mess of the polarities*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   331
lemma expand_iff [rule_format]:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   332
 "((A :: bool) = B) \<equiv> (~ A | B) & (~ B | A)"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   333
by (rule eq_reflection, auto)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   334
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   335
lemma polarity_switch [rule_format]:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   336
  "(\<not> P) = True \<Longrightarrow> P = False"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   337
  "(\<not> P) = False \<Longrightarrow> P = True"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   338
  "P = False \<Longrightarrow> (\<not> P) = True"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   339
  "P = True \<Longrightarrow> (\<not> P) = False"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   340
by auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   341
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   342
lemma solved_all_splits: "False = True \<Longrightarrow> False" by simp
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   343
ML \<open>
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   344
fun solved_all_splits_tac ctxt =
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   345
  TRY (eresolve_tac ctxt @{thms conjE} 1)
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   346
  THEN resolve_tac ctxt @{thms solved_all_splits} 1
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   347
  THEN assume_tac ctxt 1
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   348
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   349
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   350
lemma lots_of_logic_expansions_meta [rule_format]:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   351
  "(((A :: bool) = B) = True) == (((A \<longrightarrow> B) = True) & ((B \<longrightarrow> A) = True))"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   352
  "((A :: bool) = B) = False == (((~A) | B) = False) | (((~B) | A) = False)"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   353
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
   354
  "((F = G) = True) == (\<forall>x. (F x = G x)) = True"
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
   355
  "((F = G) = False) == (\<forall>x. (F x = G x)) = False"
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   356
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   357
  "(A | B) = True == (A = True) | (B = True)"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   358
  "(A & B) = False == (A = False) | (B = False)"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   359
  "(A | B) = False == (A = False) & (B = False)"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   360
  "(A & B) = True == (A = True) & (B = True)"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   361
  "(~ A) = True == A = False"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   362
  "(~ A) = False == A = True"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   363
  "~ (A = True) == A = False"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   364
  "~ (A = False) == A = True"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   365
by (rule eq_reflection, auto)+
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   366
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   367
(*this is used in extcnf_combined handler*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   368
lemma eq_neg_bool: "((A :: bool) = B) = False ==> ((~ (A | B)) | ~ ((~ A) | (~ B))) = False"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   369
by auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   370
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   371
lemma eq_pos_bool:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   372
  "((A :: bool) = B) = True ==> ((~ (A | B)) | ~ (~ A | ~ B)) = True"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   373
  "(A = B) = True \<Longrightarrow> A = True \<or> B = False"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   374
  "(A = B) = True \<Longrightarrow> A = False \<or> B = True"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   375
by auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   376
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   377
(*next formula is more versatile than
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   378
    "(F = G) = True \<Longrightarrow> \<forall>x. ((F x = G x) = True)"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   379
  since it doesn't assume that clause is singleton. After splitqtfr,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   380
  and after applying allI exhaustively to the conclusion, we can
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   381
  use the existing functions to find the "(F x = G x) = True"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   382
  disjunct in the conclusion*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   383
lemma eq_pos_func: "\<And> x. (F = G) = True \<Longrightarrow> (F x = G x) = True"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   384
by auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   385
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   386
(*make sure the conclusion consists of just "False"*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   387
lemma flip:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   388
  "((A = True) ==> False) ==> A = False"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   389
  "((A = False) ==> False) ==> A = True"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   390
by auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   391
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   392
(*FIXME try to use Drule.equal_elim_rule1 directly for this*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   393
lemma equal_elim_rule1: "(A \<equiv> B) \<Longrightarrow> A \<Longrightarrow> B" by auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   394
lemmas leo2_rules =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   395
 lots_of_logic_expansions_meta[THEN equal_elim_rule1]
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
(*FIXME is there any overlap with lots_of_logic_expansions_meta or leo2_rules?*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   398
lemma extuni_bool2 [rule_format]: "(A = B) = False \<Longrightarrow> (A = True) | (B = True)" by auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   399
lemma extuni_bool1 [rule_format]: "(A = B) = False \<Longrightarrow> (A = False) | (B = False)" by auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   400
lemma extuni_triv [rule_format]: "(A = A) = False \<Longrightarrow> R" by auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   401
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   402
(*Order (of A, B, C, D) matters*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   403
lemma dec_commut_eq [rule_format]:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   404
  "((A = B) = (C = D)) = False \<Longrightarrow> (B = C) = False | (A = D) = False"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   405
  "((A = B) = (C = D)) = False \<Longrightarrow> (B = D) = False | (A = C) = False"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   406
by auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   407
lemma dec_commut_disj [rule_format]:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   408
  "((A \<or> B) = (C \<or> D)) = False \<Longrightarrow> (B = C) = False \<or> (A = D) = False"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   409
by auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   410
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
   411
lemma extuni_func [rule_format]: "(F = G) = False \<Longrightarrow> (\<forall>X. (F X = G X)) = False" by auto
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   412
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
subsection "Emulation: tactics"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   415
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   416
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   417
(*Instantiate a variable according to the info given in the
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   418
  proof annotation. Through this we avoid having to come up
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   419
  with instantiations during reconstruction.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   420
fun bind_tac ctxt prob_name ordered_binds =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   421
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   422
    val thy = Proof_Context.theory_of ctxt
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   423
    fun term_to_string t =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   424
        Print_Mode.with_modes [""]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   425
          (fn () => Output.output (Syntax.string_of_term ctxt t)) ()
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   426
    val ordered_instances =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   427
      TPTP_Reconstruct.interpret_bindings prob_name thy ordered_binds []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   428
      |> map (snd #> term_to_string)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   429
      |> permute
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   430
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   431
    (*instantiate a list of variables, order matters*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   432
    fun instantiate_vars ctxt vars : tactic =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   433
      map (fn var =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   434
            Rule_Insts.eres_inst_tac ctxt
59780
23b67731f4f0 support 'for' fixes in rule_tac etc.;
wenzelm
parents: 59763
diff changeset
   435
             [((("x", 0), Position.none), var)] [] @{thm allE} 1)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   436
          vars
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   437
      |> EVERY
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   438
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   439
    fun instantiate_tac vars =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   440
      instantiate_vars ctxt vars
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   441
      THEN (HEADGOAL (assume_tac ctxt))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   442
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   443
    HEADGOAL (canonicalise_qtfr_order ctxt)
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   444
    THEN (REPEAT_DETERM (HEADGOAL (resolve_tac ctxt @{thms allI})))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   445
    THEN REPEAT_DETERM (HEADGOAL (nominal_inst_parametermatch_tac ctxt @{thm allE}))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   446
    (*now only the variable to instantiate should be left*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   447
    THEN FIRST (map instantiate_tac ordered_instances)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   448
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   449
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   450
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   451
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   452
(*Simplification tactics*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   453
local
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   454
  fun rew_goal_tac thms ctxt i =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   455
    rewrite_goal_tac ctxt thms i
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   456
    |> CHANGED
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   457
in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   458
  val expander_animal =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   459
    rew_goal_tac (@{thms simp_meta} @ @{thms lots_of_logic_expansions_meta})
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   460
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   461
  val simper_animal =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   462
    rew_goal_tac @{thms simp_meta}
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   463
end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   464
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   465
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   466
lemma prop_normalise [rule_format]:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   467
  "(A | B) | C == A | B | C"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   468
  "(A & B) & C == A & B & C"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   469
  "A | B == ~(~A & ~B)"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   470
  "~~ A == A"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   471
by auto
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   472
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   473
(*i.e., break_conclusion*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   474
fun flip_conclusion_tac ctxt =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   475
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   476
    val default_tac =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   477
      (TRY o CHANGED o (rewrite_goal_tac ctxt @{thms prop_normalise}))
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   478
      THEN' resolve_tac ctxt @{thms notI}
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   479
      THEN' (REPEAT_DETERM o eresolve_tac ctxt @{thms conjE})
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   480
      THEN' (TRY o (expander_animal ctxt))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   481
  in
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58941
diff changeset
   482
    default_tac ORELSE' resolve_tac ctxt @{thms flip}
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   483
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   484
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   485
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   486
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   487
subsection "Skolemisation"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   488
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   489
lemma skolemise [rule_format]:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
   490
  "\<forall> P. (\<not> (\<forall>x. P x)) \<longrightarrow> \<not> (P (SOME x. ~ P x))"
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   491
proof -
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
   492
  have "\<And> P. (\<not> (\<forall>x. P x)) \<Longrightarrow> \<not> (P (SOME x. ~ P x))"
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   493
  proof -
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   494
    fix P
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
   495
    assume ption: "\<not> (\<forall>x. P x)"
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
   496
    hence a: "\<exists>x. \<not> P x" by force
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   497
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
   498
    have hilbert : "\<And>P. (\<exists>x. P x) \<Longrightarrow> (P (SOME x. P x))"
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   499
    proof -
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   500
      fix P
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
   501
      assume "(\<exists>x. P x)"
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   502
      thus "(P (SOME x. P x))"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   503
        apply auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   504
        apply (rule someI)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   505
        apply auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   506
        done
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   507
    qed
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   508
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
   509
    from a show "\<not> P (SOME x. \<not> P x)"
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   510
    proof -
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
   511
      assume "\<exists>x. \<not> P x"
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   512
      hence "\<not> P (SOME x. \<not> P x)" by (rule hilbert)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   513
      thus ?thesis .
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   514
    qed
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   515
  qed
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   516
  thus ?thesis by blast
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   517
qed
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   518
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   519
lemma polar_skolemise [rule_format]:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
   520
  "\<forall>P. (\<forall>x. P x) = False \<longrightarrow> (P (SOME x. \<not> P x)) = False"
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   521
proof -
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
   522
  have "\<And>P. (\<forall>x. P x) = False \<Longrightarrow> (P (SOME x. \<not> P x)) = False"
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   523
  proof -
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   524
    fix P
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
   525
    assume ption: "(\<forall>x. P x) = False"
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
   526
    hence "\<not> (\<forall>x. P x)" by force
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   527
    hence "\<not> All P" by force
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   528
    hence "\<not> (P (SOME x. \<not> P x))" by (rule skolemise)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   529
    thus "(P (SOME x. \<not> P x)) = False" by force
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   530
  qed
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   531
  thus ?thesis by blast
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   532
qed
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   533
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   534
lemma leo2_skolemise [rule_format]:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
   535
  "\<forall>P sk. (\<forall>x. P x) = False \<longrightarrow> (sk = (SOME x. \<not> P x)) \<longrightarrow> (P sk) = False"
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   536
by (clarify, rule polar_skolemise)
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
lemma lift_forall [rule_format]:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
   539
  "\<And>x. (\<forall>x. A x) = True \<Longrightarrow> (A x) = True"
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
   540
  "\<And>x. (\<exists>x. A x) = False \<Longrightarrow> (A x) = False"
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   541
by auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   542
lemma lift_exists [rule_format]:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   543
  "\<lbrakk>(All P) = False; sk = (SOME x. \<not> P x)\<rbrakk> \<Longrightarrow> P sk = False"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   544
  "\<lbrakk>(Ex P) = True; sk = (SOME x. P x)\<rbrakk> \<Longrightarrow> P sk = True"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   545
apply (drule polar_skolemise, simp)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   546
apply (simp, drule someI_ex, simp)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   547
done
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   548
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   549
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   550
(*FIXME LHS should be constant. Currently allow variables for testing. Probably should still allow Vars (but not Frees) since they'll act as intermediate values*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   551
fun conc_is_skolem_def t =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   552
  case t of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   553
      Const (\<^const_name>\<open>HOL.eq\<close>, _) $ t' $ (Const (\<^const_name>\<open>Hilbert_Choice.Eps\<close>, _) $ _) =>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   554
      let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   555
        val (h, args) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   556
          strip_comb t'
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   557
          |> apfst (strip_abs #> snd #> strip_comb #> fst)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   558
        val get_const_name = dest_Const #> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   559
        val h_property =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   560
          is_Free h orelse
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   561
          is_Var h orelse
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   562
          (is_Const h
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   563
           andalso (get_const_name h <> get_const_name \<^term>\<open>HOL.Ex\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   564
           andalso (get_const_name h <> get_const_name \<^term>\<open>HOL.All\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   565
           andalso (h <> \<^term>\<open>Hilbert_Choice.Eps\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   566
           andalso (h <> \<^term>\<open>HOL.conj\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   567
           andalso (h <> \<^term>\<open>HOL.disj\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   568
           andalso (h <> \<^term>\<open>HOL.eq\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   569
           andalso (h <> \<^term>\<open>HOL.implies\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   570
           andalso (h <> \<^term>\<open>HOL.The\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   571
           andalso (h <> \<^term>\<open>HOL.Ex1\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   572
           andalso (h <> \<^term>\<open>HOL.Not\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   573
           andalso (h <> \<^term>\<open>HOL.iff\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   574
           andalso (h <> \<^term>\<open>HOL.not_equal\<close>))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   575
        val args_property =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   576
          fold (fn t => fn b =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   577
           b andalso is_Free t) args true
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   578
      in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   579
        h_property andalso args_property
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   580
      end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   581
    | _ => false
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   582
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   583
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   584
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   585
(*Hack used to detect if a Skolem definition, with an LHS Var, has had the LHS instantiated into an unacceptable term.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   586
fun conc_is_bad_skolem_def t =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   587
  case t of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   588
      Const (\<^const_name>\<open>HOL.eq\<close>, _) $ t' $ (Const (\<^const_name>\<open>Hilbert_Choice.Eps\<close>, _) $ _) =>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   589
      let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   590
        val (h, args) = strip_comb t'
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   591
        val get_const_name = dest_Const #> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   592
        val const_h_test =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   593
          if is_Const h then
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   594
            (get_const_name h = get_const_name \<^term>\<open>HOL.Ex\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   595
             orelse (get_const_name h = get_const_name \<^term>\<open>HOL.All\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   596
             orelse (h = \<^term>\<open>Hilbert_Choice.Eps\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   597
             orelse (h = \<^term>\<open>HOL.conj\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   598
             orelse (h = \<^term>\<open>HOL.disj\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   599
             orelse (h = \<^term>\<open>HOL.eq\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   600
             orelse (h = \<^term>\<open>HOL.implies\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   601
             orelse (h = \<^term>\<open>HOL.The\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   602
             orelse (h = \<^term>\<open>HOL.Ex1\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   603
             orelse (h = \<^term>\<open>HOL.Not\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   604
             orelse (h = \<^term>\<open>HOL.iff\<close>)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   605
             orelse (h = \<^term>\<open>HOL.not_equal\<close>)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   606
          else true
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   607
        val h_property =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   608
          not (is_Free h) andalso
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   609
          not (is_Var h) andalso
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   610
          const_h_test
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   611
        val args_property =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   612
          fold (fn t => fn b =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   613
           b andalso is_Free t) args true
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   614
      in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   615
        h_property andalso args_property
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   616
      end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   617
    | _ => false
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   618
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   619
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   620
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   621
fun get_skolem_conc t =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   622
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   623
    val t' =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   624
      strip_top_all_vars [] t
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   625
      |> snd
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   626
      |> try_dest_Trueprop
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   627
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   628
    case t' of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   629
        Const (\<^const_name>\<open>HOL.eq\<close>, _) $ t' $ (Const (\<^const_name>\<open>Hilbert_Choice.Eps\<close>, _) $ _) => SOME t'
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   630
      | _ => NONE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   631
  end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   632
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   633
fun get_skolem_conc_const t =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   634
  lift_option
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   635
   (fn t' =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   636
     head_of t'
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   637
     |> strip_abs_body
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   638
     |> head_of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   639
     |> dest_Const)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   640
   (get_skolem_conc t)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   641
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   642
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   643
(*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   644
Technique for handling quantifiers:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   645
  Principles:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   646
  * allE should always match with a !!
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   647
  * exE should match with a constant,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   648
     or bind a fresh !! -- currently not doing the latter since it never seems to arised in normal Leo2 proofs.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   649
*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   650
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   651
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   652
fun forall_neg_tac candidate_consts ctxt i = fn st =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   653
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   654
    val gls =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59533
diff changeset
   655
      Thm.prop_of st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   656
      |> Logic.strip_horn
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   657
      |> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   658
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   659
    val parameters =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   660
      if null gls then ""
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   661
      else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   662
        rpair (i - 1) gls
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   663
        |> uncurry nth
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   664
        |> strip_top_all_vars []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   665
        |> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   666
        |> map fst (*just get the parameter names*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   667
        |> (fn l =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   668
              if null l then ""
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   669
              else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   670
                space_implode " " l
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   671
                |> pair " "
67405
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67404
diff changeset
   672
                |> (op ^))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   673
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   674
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   675
    if null gls orelse null candidate_consts then no_tac st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   676
    else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   677
      let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   678
        fun instantiate const_name =
59780
23b67731f4f0 support 'for' fixes in rule_tac etc.;
wenzelm
parents: 59763
diff changeset
   679
          Rule_Insts.dres_inst_tac ctxt [((("sk", 0), Position.none), const_name ^ parameters)] []
59755
f8d164ab0dc1 more position information;
wenzelm
parents: 59639
diff changeset
   680
            @{thm leo2_skolemise}
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   681
        val attempts = map instantiate candidate_consts
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   682
      in
67405
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67404
diff changeset
   683
        (fold (curry (op APPEND')) attempts (K no_tac)) i st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   684
      end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   685
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   686
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   687
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   688
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   689
exception SKOLEM_DEF of term (*The tactic wasn't pointed at a skolem definition*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   690
exception NO_SKOLEM_DEF of (*skolem const name*)string * Binding.binding * term (*The tactic could not find a skolem definition in the theory*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   691
fun absorb_skolem_def ctxt prob_name_opt i = fn st =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   692
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   693
    val thy = Proof_Context.theory_of ctxt
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   694
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   695
    val gls =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59533
diff changeset
   696
      Thm.prop_of st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   697
      |> Logic.strip_horn
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   698
      |> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   699
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   700
    val conclusion =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   701
      if null gls then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   702
        (*this should never be thrown*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   703
        raise NO_GOALS
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   704
      else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   705
        rpair (i - 1) gls
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   706
        |> uncurry nth
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   707
        |> strip_top_all_vars []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   708
        |> snd
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   709
        |> Logic.strip_horn
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   710
        |> snd
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   711
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   712
    fun skolem_const_info_of t =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   713
      case t of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   714
          Const (\<^const_name>\<open>HOL.Trueprop\<close>, _) $ (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ t' $ (Const (\<^const_name>\<open>Hilbert_Choice.Eps\<close>, _) $ _)) =>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   715
          head_of t'
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   716
          |> strip_abs_body (*since in general might have a skolem term, so we want to rip out the prefixing lambdas to get to the constant (which should be at head position)*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   717
          |> head_of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   718
          |> dest_Const
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   719
        | _ => raise SKOLEM_DEF t
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
    val const_name =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   722
      skolem_const_info_of conclusion
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   723
      |> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   724
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   725
    val def_name = const_name ^ "_def"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   726
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   727
    val bnd_def = (*FIXME consts*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   728
      const_name
56220
4c43a2881b25 more explicit Long_Name operations (NB: analyzing qualifiers is inherently fragile);
wenzelm
parents: 55597
diff changeset
   729
      |> Long_Name.implode o tl o Long_Name.explode (*FIXME hack to drop theory-name prefix*)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   730
      |> Binding.qualified_name
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   731
      |> Binding.suffix_name "_def"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   732
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   733
    val bnd_name =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   734
      case prob_name_opt of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   735
          NONE => bnd_def
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   736
        | SOME prob_name =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   737
(*            Binding.qualify false
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   738
             (TPTP_Problem_Name.mangle_problem_name prob_name)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   739
*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   740
             bnd_def
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   741
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   742
    val thm =
59882
ada832308efe tuned -- avoid exotic Name_Space.defined_entry;
wenzelm
parents: 59858
diff changeset
   743
      (case try (Thm.axiom thy) def_name of
ada832308efe tuned -- avoid exotic Name_Space.defined_entry;
wenzelm
parents: 59858
diff changeset
   744
        SOME thm => thm
ada832308efe tuned -- avoid exotic Name_Space.defined_entry;
wenzelm
parents: 59858
diff changeset
   745
      | NONE =>
ada832308efe tuned -- avoid exotic Name_Space.defined_entry;
wenzelm
parents: 59858
diff changeset
   746
          if is_none prob_name_opt then
ada832308efe tuned -- avoid exotic Name_Space.defined_entry;
wenzelm
parents: 59858
diff changeset
   747
            (*This mode is for testing, so we can be a bit
ada832308efe tuned -- avoid exotic Name_Space.defined_entry;
wenzelm
parents: 59858
diff changeset
   748
              looser with theories*)
ada832308efe tuned -- avoid exotic Name_Space.defined_entry;
wenzelm
parents: 59858
diff changeset
   749
            (* FIXME bad theory context!? *)
ada832308efe tuned -- avoid exotic Name_Space.defined_entry;
wenzelm
parents: 59858
diff changeset
   750
            Thm.add_axiom_global (bnd_name, conclusion) thy
ada832308efe tuned -- avoid exotic Name_Space.defined_entry;
wenzelm
parents: 59858
diff changeset
   751
            |> fst |> snd
ada832308efe tuned -- avoid exotic Name_Space.defined_entry;
wenzelm
parents: 59858
diff changeset
   752
          else
ada832308efe tuned -- avoid exotic Name_Space.defined_entry;
wenzelm
parents: 59858
diff changeset
   753
            raise (NO_SKOLEM_DEF (def_name, bnd_name, conclusion)))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   754
  in
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   755
    resolve_tac ctxt [Drule.export_without_context thm] i st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   756
  end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   757
  handle SKOLEM_DEF _ => no_tac st
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   758
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   759
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   760
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   761
(*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   762
In current system, there should only be 2 subgoals: the one where
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   763
the skolem definition is being built (with a Var in the LHS), and the other subgoal using Var.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   764
*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   765
(*arity must be greater than 0. if arity=0 then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   766
  there's no need to use this expensive matching.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   767
fun find_skolem_term ctxt consts_candidate arity = fn st =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   768
  let
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   769
    val _ = \<^assert> (arity > 0)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   770
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   771
    val gls =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59533
diff changeset
   772
      Thm.prop_of st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   773
      |> Logic.strip_horn
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   774
      |> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   775
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   776
    (*extract the conclusion of each subgoal*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   777
    val conclusions =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   778
      if null gls then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   779
        raise NO_GOALS
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   780
      else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   781
        map (strip_top_all_vars [] #> snd #> Logic.strip_horn #> snd) gls
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   782
        (*Remove skolem-definition conclusion, to avoid wasting time analysing it*)
58412
f65f11f4854c more standard Isabelle/ML operations;
wenzelm
parents: 58411
diff changeset
   783
        |> filter (try_dest_Trueprop #> conc_is_skolem_def #> not)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   784
        (*There should only be a single goal*) (*FIXME this might not always be the case, in practice*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   785
        (* |> tap (fn x => @{assert} (is_some (try the_single x))) *)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   786
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   787
    (*look for subterms headed by a skolem constant, and whose
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   788
      arguments are all parameter Vars*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   789
    fun get_skolem_terms args (acc : term list) t =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   790
      case t of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   791
          (c as Const _) $ (v as Free _) =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   792
            if c = consts_candidate andalso
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   793
             arity = length args + 1 then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   794
              (list_comb (c, v :: args)) :: acc
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   795
            else acc
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   796
        | t1 $ (v as Free _) =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   797
            get_skolem_terms (v :: args) acc t1 @
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   798
             get_skolem_terms [] acc t1
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   799
        | t1 $ t2 =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   800
            get_skolem_terms [] acc t1 @
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   801
             get_skolem_terms [] acc t2
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   802
        | Abs (_, _, t') => get_skolem_terms [] acc t'
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   803
        | _ => acc
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   804
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   805
    map (strip_top_All_vars #> snd) conclusions
58411
wenzelm
parents: 57773
diff changeset
   806
    |> maps (get_skolem_terms [] [])
67405
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67404
diff changeset
   807
    |> distinct (op =)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   808
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   809
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   810
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   811
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   812
fun instantiate_skols ctxt consts_candidates i = fn st =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   813
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   814
    val gls =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59533
diff changeset
   815
      Thm.prop_of st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   816
      |> Logic.strip_horn
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   817
      |> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   818
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   819
    val (params, conclusion) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   820
      if null gls then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   821
        raise NO_GOALS
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   822
      else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   823
        rpair (i - 1) gls
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   824
        |> uncurry nth
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   825
        |> strip_top_all_vars []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   826
        |> apsnd (Logic.strip_horn #> snd)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   827
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   828
    fun skolem_const_info_of t =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   829
      case t of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   830
          Const (\<^const_name>\<open>HOL.Trueprop\<close>, _) $ (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ lhs $ (Const (\<^const_name>\<open>Hilbert_Choice.Eps\<close>, _) $ rhs)) =>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   831
          let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   832
            (*the parameters we will concern ourselves with*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   833
            val params' =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   834
              Term.add_frees lhs []
67405
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67404
diff changeset
   835
              |> distinct (op =)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   836
            (*check to make sure that params' <= params*)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   837
            val _ = \<^assert> (forall (member (op =) params) params')
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   838
            val skolem_const_ty =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   839
              let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   840
                val (skolem_const_prety, no_params) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   841
                  Term.strip_comb lhs
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   842
                  |> apfst (dest_Var #> snd) (*head of lhs consists of a logical variable. we just want its type.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   843
                  |> apsnd length
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   844
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   845
                val _ = \<^assert> (length params = no_params)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   846
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   847
                (*get value type of a function type after n arguments have been supplied*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   848
                fun get_val_ty n ty =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   849
                  if n = 0 then ty
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   850
                  else get_val_ty (n - 1) (dest_funT ty |> snd)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   851
              in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   852
                get_val_ty no_params skolem_const_prety
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   853
              end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   854
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   855
          in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   856
            (skolem_const_ty, params')
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   857
          end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   858
        | _ => raise (SKOLEM_DEF t)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   859
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   860
(*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   861
find skolem const candidates which, after applying distinct members of params' we end up with, give us something of type skolem_const_ty.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   862
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   863
given a candidate's type, skolem_const_ty, and params', we get some pemutations of params' (i.e. the order in which they can be given to the candidate in order to get skolem_const_ty). If the list of permutations is empty, then we cannot use that candidate.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   864
*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   865
(*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   866
only returns a single matching -- since terms are linear, and variable arguments are Vars, order shouldn't matter, so we can ignore permutations.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   867
doesn't work with polymorphism (for which we'd need to use type unification) -- this is OK since no terms should be polymorphic, since Leo2 proofs aren't.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   868
*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   869
    fun use_candidate target_ty params acc cur_ty =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   870
      if null params then
58941
f09dd46352ba more direct type equality;
wenzelm
parents: 58412
diff changeset
   871
        if cur_ty = target_ty then
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   872
          SOME (rev acc)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   873
        else NONE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   874
      else
78727
1b052426a2b7 avoid accidental 'handle' of interrupts;
wenzelm
parents: 77879
diff changeset
   875
        \<^try>\<open>
1b052426a2b7 avoid accidental 'handle' of interrupts;
wenzelm
parents: 77879
diff changeset
   876
          let
1b052426a2b7 avoid accidental 'handle' of interrupts;
wenzelm
parents: 77879
diff changeset
   877
            val (arg_ty, val_ty) = Term.dest_funT cur_ty
1b052426a2b7 avoid accidental 'handle' of interrupts;
wenzelm
parents: 77879
diff changeset
   878
            (*now find a param of type arg_ty*)
1b052426a2b7 avoid accidental 'handle' of interrupts;
wenzelm
parents: 77879
diff changeset
   879
            val (candidate_param, params') =
1b052426a2b7 avoid accidental 'handle' of interrupts;
wenzelm
parents: 77879
diff changeset
   880
              find_and_remove (snd #> pair arg_ty #> (op =)) params
1b052426a2b7 avoid accidental 'handle' of interrupts;
wenzelm
parents: 77879
diff changeset
   881
          in
1b052426a2b7 avoid accidental 'handle' of interrupts;
wenzelm
parents: 77879
diff changeset
   882
            use_candidate target_ty params' (candidate_param :: acc) val_ty
1b052426a2b7 avoid accidental 'handle' of interrupts;
wenzelm
parents: 77879
diff changeset
   883
          end
1b052426a2b7 avoid accidental 'handle' of interrupts;
wenzelm
parents: 77879
diff changeset
   884
          catch TYPE ("dest_funT", _, _) => NONE  (* FIXME fragile *)
62243
dd22d2423047 clarified exception handling;
wenzelm
parents: 60754
diff changeset
   885
             | _ => NONE  (* FIXME avoid catch-all handler *)
78727
1b052426a2b7 avoid accidental 'handle' of interrupts;
wenzelm
parents: 77879
diff changeset
   886
        \<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   887
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   888
    val (skolem_const_ty, params') = skolem_const_info_of conclusion
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   889
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   890
(*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   891
For each candidate, build a term and pass it to Thm.instantiate, whic in turn is chained with PRIMITIVE to give us this_tactic.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   892
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   893
Big picture:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   894
  we run the following:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   895
    drule leo2_skolemise THEN' this_tactic
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   896
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   897
NOTE: remember to APPEND' instead of ORELSE' the two tactics relating to skolemisation
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   898
*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   899
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   900
    val filtered_candidates =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   901
      map (dest_Const
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   902
           #> snd
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   903
           #> use_candidate skolem_const_ty params' [])
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   904
       consts_candidates (* prefiltered_candidates *)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   905
      |> pair consts_candidates (* prefiltered_candidates *)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   906
      |> ListPair.zip
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   907
      |> filter (snd #> is_none #> not)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   908
      |> map (apsnd the)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   909
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   910
    val skolem_terms =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   911
      let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   912
        fun make_result_t (t, args) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   913
          (* list_comb (t, map Free args) *)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   914
          if length args > 0 then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   915
            hd (find_skolem_term ctxt t (length args) st)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   916
          else t
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   917
      in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   918
        map make_result_t filtered_candidates
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   919
      end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   920
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   921
    (*prefix a skolem term with bindings for the parameters*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   922
    (* val contextualise = fold absdummy (map snd params) *)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   923
    val contextualise = fold absfree params
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   924
59639
f596ed647018 clarified context;
wenzelm
parents: 59621
diff changeset
   925
    val skolem_cts = map (contextualise #> Thm.cterm_of ctxt) skolem_terms
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   926
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   927
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   928
(*now the instantiation code*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   929
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   930
    (*there should only be one Var -- that is from the previous application of drule leo2_skolemise. We look for it at the head position in some equation at a conclusion of a subgoal.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   931
    val var_opt =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   932
      let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   933
        val pre_var =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   934
          gls
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   935
          |> map
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   936
              (strip_top_all_vars [] #> snd #>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   937
               Logic.strip_horn #> snd #>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   938
               get_skolem_conc)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   939
          |> switch (fold (fn x => fn l => if is_some x then the x :: l else l)) []
58411
wenzelm
parents: 57773
diff changeset
   940
          |> maps (switch Term.add_vars [])
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   941
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   942
        fun make_var pre_var =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   943
          the_single pre_var
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   944
          |> Var
59639
f596ed647018 clarified context;
wenzelm
parents: 59621
diff changeset
   945
          |> Thm.cterm_of ctxt
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   946
          |> SOME
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   947
      in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   948
        if null pre_var then NONE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   949
        else make_var pre_var
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   950
     end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   951
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   952
    fun instantiate_tac from to =
77879
wenzelm
parents: 74282
diff changeset
   953
      PRIMITIVE (Thm.instantiate (TVars.empty, Vars.make1 (from, to)))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   954
59760
wenzelm
parents: 59755
diff changeset
   955
    val tactic =
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   956
      if is_none var_opt then no_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   957
      else
67405
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67404
diff changeset
   958
        fold (curry (op APPEND))
60642
48dd1cefb4ae simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents: 60319
diff changeset
   959
          (map (instantiate_tac (dest_Var (Thm.term_of (the var_opt)))) skolem_cts) no_tac
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   960
  in
59760
wenzelm
parents: 59755
diff changeset
   961
    tactic st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   962
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   963
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   964
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   965
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   966
fun new_skolem_tac ctxt consts_candidates =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   967
  let
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   968
    fun tac thm =
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   969
      dresolve_tac ctxt [thm]
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   970
      THEN' instantiate_skols ctxt consts_candidates
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   971
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   972
    if null consts_candidates then K no_tac
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   973
    else FIRST' (map tac @{thms lift_exists})
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   974
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   975
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   976
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   977
(*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   978
need a tactic to expand "? x . P" to "~ ! x. ~ P"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   979
*)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   980
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   981
fun ex_expander_tac ctxt i =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   982
   let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   983
     val simpset =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   984
       empty_simpset ctxt (*NOTE for some reason, Bind exception gets raised if ctxt's simpset isn't emptied*)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
   985
       |> Simplifier.add_simp @{lemma "Ex P == (\<not> (\<forall>x. \<not> P x))" by auto}
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   986
   in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   987
     CHANGED (asm_full_simp_tac simpset i)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   988
   end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   989
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   990
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   991
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   992
subsubsection "extuni_dec"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   993
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
   994
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   995
(*n-ary decomposition. Code is based on the n-ary arg_cong generator*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   996
fun extuni_dec_n ctxt arity =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   997
  let
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
   998
    val _ = \<^assert> (arity > 0)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
   999
    val is =
67404
e128ab544996 proper infix;
wenzelm
parents: 67399
diff changeset
  1000
      1 upto arity
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1001
      |> map Int.toString
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1002
    val arg_tys = map (fn i => TFree ("arg" ^ i ^ "_ty", \<^sort>\<open>type\<close>)) is
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1003
    val res_ty = TFree ("res" ^ "_ty", \<^sort>\<open>type\<close>)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1004
    val f_ty = arg_tys ---> res_ty
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1005
    val f = Free ("f", f_ty)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1006
    val xs = map (fn i =>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1007
      Free ("x" ^ i, TFree ("arg" ^ i ^ "_ty", \<^sort>\<open>type\<close>))) is
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1008
    (*FIXME DRY principle*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1009
    val ys = map (fn i =>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1010
      Free ("y" ^ i, TFree ("arg" ^ i ^ "_ty", \<^sort>\<open>type\<close>))) is
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1011
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1012
    val hyp_lhs = list_comb (f, xs)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1013
    val hyp_rhs = list_comb (f, ys)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1014
    val hyp_eq =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1015
      HOLogic.eq_const res_ty $ hyp_lhs $ hyp_rhs
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1016
    val hyp =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1017
      HOLogic.eq_const HOLogic.boolT $ hyp_eq $ \<^term>\<open>False\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1018
      |> HOLogic.mk_Trueprop
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1019
    fun conc_eq i =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1020
      let
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1021
        val ty = TFree ("arg" ^ i ^ "_ty", \<^sort>\<open>type\<close>)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1022
        val x = Free ("x" ^ i, ty)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1023
        val y = Free ("y" ^ i, ty)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1024
        val eq = HOLogic.eq_const ty $ x $ y
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1025
      in
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1026
        HOLogic.eq_const HOLogic.boolT $ eq $ \<^term>\<open>False\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1027
      end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1028
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1029
    val conc_disjs = map conc_eq is
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1030
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1031
    val conc =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1032
      if length conc_disjs = 1 then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1033
        the_single conc_disjs
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1034
      else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1035
        fold
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1036
         (fn t => fn t_conc => HOLogic.mk_disj (t_conc, t))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1037
         (tl conc_disjs) (hd conc_disjs)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1038
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1039
    val t =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1040
      Logic.mk_implies (hyp, HOLogic.mk_Trueprop conc)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1041
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1042
    Goal.prove ctxt [] [] t (fn _ => auto_tac ctxt)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1043
    |> Drule.export_without_context
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1044
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1045
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1046
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1047
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1048
(*Determine the arity of a function which the "dec"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1049
  unification rule is about to be applied.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1050
  NOTE:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1051
    * Assumes that there is a single hypothesis
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1052
*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1053
fun find_dec_arity i = fn st =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1054
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1055
    val gls =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59533
diff changeset
  1056
      Thm.prop_of st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1057
      |> Logic.strip_horn
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1058
      |> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1059
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1060
    if null gls then raise NO_GOALS
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1061
    else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1062
      let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1063
        val (params, (literal, conc_clause)) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1064
          rpair (i - 1) gls
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1065
          |> uncurry nth
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1066
          |> strip_top_all_vars []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1067
          |> apsnd Logic.strip_horn
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1068
          |> apsnd (apfst the_single)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1069
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1070
        val get_ty =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1071
          HOLogic.dest_Trueprop
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1072
          #> strip_top_All_vars
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1073
          #> snd
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1074
          #> HOLogic.dest_eq (*polarity's "="*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1075
          #> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1076
          #> HOLogic.dest_eq (*the unification constraint's "="*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1077
          #> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1078
          #> head_of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1079
          #> dest_Const
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1080
          #> snd
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1081
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1082
       fun arity_of ty =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1083
         let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1084
           val (_, res_ty) = dest_funT ty
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1085
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1086
         in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1087
           1 + arity_of res_ty
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1088
         end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1089
         handle (TYPE ("dest_funT", _, _)) => 0
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1090
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1091
      in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1092
        arity_of (get_ty literal)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1093
      end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1094
  end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1095
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1096
(*given an inference, it returns the parameters (i.e., we've already matched the leading & shared quantification in the hypothesis & conclusion clauses), and the "raw" inference*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1097
fun breakdown_inference i = fn st =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1098
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1099
    val gls =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59533
diff changeset
  1100
      Thm.prop_of st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1101
      |> Logic.strip_horn
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1102
      |> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1103
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1104
    if null gls then raise NO_GOALS
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1105
    else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1106
      rpair (i - 1) gls
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1107
      |> uncurry nth
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1108
      |> strip_top_all_vars []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1109
  end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1110
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1111
(*build a custom elimination rule for extuni_dec, and instantiate it to match a specific subgoal*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1112
fun extuni_dec_elim_rule ctxt arity i = fn st =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1113
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1114
    val rule = extuni_dec_n ctxt arity
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1115
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1116
    val rule_hyp =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59533
diff changeset
  1117
      Thm.prop_of rule
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1118
      |> Logic.dest_implies
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1119
      |> fst (*assuming that rule has single hypothesis*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1120
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1121
    (*having run break_hypothesis earlier, we know that the hypothesis
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1122
      now consists of a single literal. We can (and should)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1123
      disregard the conclusion, since it hasn't been "broken",
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1124
      and it might include some unwanted literals -- the latter
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1125
      could cause "diff" to fail (since they won't agree with the
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1126
      rule we have generated.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1127
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1128
    val inference_hyp =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1129
      snd (breakdown_inference i st)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1130
      |> Logic.dest_implies
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1131
      |> fst (*assuming that inference has single hypothesis,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1132
               as explained above.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1133
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1134
    TPTP_Reconstruct_Library.diff_and_instantiate ctxt rule rule_hyp inference_hyp
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1135
  end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1136
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1137
fun extuni_dec_tac ctxt i = fn st =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1138
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1139
    val arity = find_dec_arity i st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1140
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1141
    fun elim_tac i st =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1142
      let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1143
        val rule =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1144
          extuni_dec_elim_rule ctxt arity i st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1145
          (*in case we itroduced free variables during
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1146
            instantiation, we generalise the rule to make
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1147
            those free variables into logical variables.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1148
          |> Thm.forall_intr_frees
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1149
          |> Drule.export_without_context
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1150
      in dresolve_tac ctxt [rule] i st end
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1151
      handle NO_GOALS => no_tac st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1152
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1153
    fun closure tac =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1154
     (*batter fails if there's no toplevel disjunction in the
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1155
       hypothesis, so we also try atac*)
59533
e9ffe89a20a4 proper context;
wenzelm
parents: 59498
diff changeset
  1156
      SOLVE o (tac THEN' (batter_tac ctxt ORELSE' assume_tac ctxt))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1157
    val search_tac =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1158
      ASAP
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1159
        (resolve_tac ctxt @{thms disjI1} APPEND' resolve_tac ctxt @{thms disjI2})
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1160
        (FIRST' (map closure
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58941
diff changeset
  1161
                  [dresolve_tac ctxt @{thms dec_commut_eq},
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1162
                   dresolve_tac ctxt @{thms dec_commut_disj},
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1163
                   elim_tac]))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1164
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1165
    (CHANGED o search_tac) i st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1166
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1167
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1168
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1169
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1170
subsubsection "standard_cnf"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1171
(*Given a standard_cnf inference, normalise it
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1172
     e.g. ((A & B) & C \<longrightarrow> D & E \<longrightarrow> F \<longrightarrow> G) = False
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1173
     is changed to
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1174
          (A & B & C & D & E & F \<longrightarrow> G) = False
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1175
 then custom-build a metatheorem which validates this:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1176
          (A & B & C & D & E & F \<longrightarrow> G) = False
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1177
       -------------------------------------------
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1178
          (A = True) & (B = True) & (C = True) &
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1179
          (D = True) & (E = True) & (F = True) & (G = False)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1180
 and apply this metatheorem.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1181
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1182
There aren't any "positive" standard_cnfs in Leo2's calculus:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1183
  e.g.,  "(A \<longrightarrow> B) = True \<Longrightarrow> A = False | (A = True & B = True)"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1184
since "standard_cnf" seems to be applied at the preprocessing
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1185
stage, together with splitting.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1186
*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1187
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1188
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1189
(*Conjunctive counterparts to Term.disjuncts_aux and Term.disjuncts*)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1190
fun conjuncts_aux (Const (\<^const_name>\<open>HOL.conj\<close>, _) $ t $ t') conjs =
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1191
     conjuncts_aux t (conjuncts_aux t' conjs)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1192
  | conjuncts_aux t conjs = t :: conjs
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1193
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1194
fun conjuncts t = conjuncts_aux t []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1195
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1196
(*HOL equivalent of Logic.strip_horn*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1197
local
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1198
  fun imp_strip_horn' acc (Const (\<^const_name>\<open>HOL.implies\<close>, _) $ A $ B) =
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1199
        imp_strip_horn' (A :: acc) B
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1200
    | imp_strip_horn' acc t = (acc, t)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1201
in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1202
  fun imp_strip_horn t =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1203
    imp_strip_horn' [] t
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1204
    |> apfst rev
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1205
end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1206
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1207
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1208
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1209
(*Returns whether the antecedents are separated by conjunctions
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1210
  or implications; the number of antecedents; and the polarity
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1211
  of the original clause -- I think this will always be "false".*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1212
fun standard_cnf_type ctxt i : thm -> (TPTP_Reconstruct.formula_kind * int * bool) option = fn st =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1213
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1214
    val gls =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59533
diff changeset
  1215
      Thm.prop_of st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1216
      |> Logic.strip_horn
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1217
      |> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1218
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1219
    val hypos =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1220
      if null gls then raise NO_GOALS
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1221
      else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1222
        rpair (i - 1) gls
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1223
        |> uncurry nth
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1224
        |> TPTP_Reconstruct.strip_top_all_vars []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1225
        |> snd
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1226
        |> Logic.strip_horn
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1227
        |> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1228
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1229
    (*hypothesis clause should be singleton*)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1230
    val _ = \<^assert> (length hypos = 1)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1231
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1232
    val (t, pol) = the_single hypos
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1233
      |> try_dest_Trueprop
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1234
      |> TPTP_Reconstruct.strip_top_All_vars
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1235
      |> snd
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1236
      |> TPTP_Reconstruct.remove_polarity true
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1237
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1238
    (*literal is negative*)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1239
    val _ = \<^assert> (not pol)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1240
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1241
    val (antes, conc) = imp_strip_horn t
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1242
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1243
    val (ante_type, antes') =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1244
      if length antes = 1 then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1245
        let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1246
          val conjunctive_antes =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1247
            the_single antes
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1248
            |> conjuncts
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1249
        in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1250
          if length conjunctive_antes > 1 then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1251
            (TPTP_Reconstruct.Conjunctive NONE,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1252
             conjunctive_antes)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1253
          else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1254
            (TPTP_Reconstruct.Implicational NONE,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1255
             antes)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1256
        end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1257
      else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1258
        (TPTP_Reconstruct.Implicational NONE,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1259
         antes)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1260
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1261
    if null antes then NONE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1262
    else SOME (ante_type, length antes', pol)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1263
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1264
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1265
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1266
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1267
(*Given a certain standard_cnf type, build a metatheorem that would
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1268
  validate it*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1269
fun mk_standard_cnf ctxt kind arity =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1270
  let
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1271
    val _ = \<^assert> (arity > 0)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1272
    val vars =
67404
e128ab544996 proper infix;
wenzelm
parents: 67399
diff changeset
  1273
      1 upto (arity + 1)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1274
      |> map (fn i => Free ("x" ^ Int.toString i, HOLogic.boolT))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1275
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1276
    val consequent = hd vars
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1277
    val antecedents = tl vars
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1278
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1279
    val conc =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1280
      fold
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1281
       (curry HOLogic.mk_conj)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1282
       (map (fn var => HOLogic.mk_eq (var, \<^term>\<open>True\<close>)) antecedents)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1283
       (HOLogic.mk_eq (consequent, \<^term>\<open>False\<close>))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1284
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1285
    val pre_hyp =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1286
      case kind of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1287
          TPTP_Reconstruct.Conjunctive NONE =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1288
            curry HOLogic.mk_imp
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1289
             (if length antecedents = 1 then the_single antecedents
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1290
              else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1291
                fold (curry HOLogic.mk_conj) (tl antecedents) (hd antecedents))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1292
             (hd vars)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1293
        | TPTP_Reconstruct.Implicational NONE =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1294
            fold (curry HOLogic.mk_imp) antecedents consequent
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1295
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1296
    val hyp = HOLogic.mk_eq (pre_hyp, \<^term>\<open>False\<close>)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1297
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1298
    val t =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1299
      Logic.mk_implies (HOLogic.mk_Trueprop  hyp, HOLogic.mk_Trueprop conc)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1300
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1301
    Goal.prove ctxt [] [] t (fn _ => HEADGOAL (blast_tac ctxt))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1302
    |> Drule.export_without_context
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1303
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1304
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1305
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1306
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1307
(*Applies a d-tactic, then breaks it up conjunctively.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1308
  This can be used to transform subgoals as follows:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1309
     (A \<longrightarrow> B) = False  \<Longrightarrow> R
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1310
              |
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1311
              v
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1312
  \<lbrakk>A = True; B = False\<rbrakk> \<Longrightarrow> R
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1313
*)
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1314
fun weak_conj_tac ctxt drule =
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1315
  dresolve_tac ctxt [drule] THEN'
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1316
  (REPEAT_DETERM o eresolve_tac ctxt @{thms conjE})
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1317
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1318
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1319
ML \<open>
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1320
fun uncurry_lit_neg_tac ctxt =
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1321
  REPEAT_DETERM o
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1322
    dresolve_tac ctxt [@{lemma "(A \<longrightarrow> B \<longrightarrow> C) = False \<Longrightarrow> (A & B \<longrightarrow> C) = False" by auto}]
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1323
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1324
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1325
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1326
fun standard_cnf_tac ctxt i = fn st =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1327
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1328
    fun core_tactic i = fn st =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1329
      case standard_cnf_type ctxt i st of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1330
          NONE => no_tac st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1331
        | SOME (kind, arity, _) =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1332
            let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1333
              val rule = mk_standard_cnf ctxt kind arity;
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1334
            in
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1335
              (weak_conj_tac ctxt rule THEN' assume_tac ctxt) i st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1336
            end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1337
  in
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1338
    (uncurry_lit_neg_tac ctxt
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1339
     THEN' TPTP_Reconstruct_Library.reassociate_conjs_tac ctxt
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1340
     THEN' core_tactic) i st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1341
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1342
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1343
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1344
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1345
subsubsection "Emulator prep"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1346
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1347
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1348
datatype cleanup_feature =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1349
    RemoveHypothesesFromSkolemDefs
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1350
  | RemoveDuplicates
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1351
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1352
datatype loop_feature =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1353
    Close_Branch
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1354
  | ConjI
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1355
  | King_Cong
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1356
  | Break_Hypotheses
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1357
  | Donkey_Cong (*simper_animal + ex_expander_tac*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1358
  | RemoveRedundantQuantifications
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1359
  | Assumption
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1360
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1361
  (*Closely based on Leo2 calculus*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1362
  | Existential_Free
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1363
  | Existential_Var
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1364
  | Universal
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1365
  | Not_pos
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1366
  | Not_neg
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1367
  | Or_pos
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1368
  | Or_neg
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1369
  | Equal_pos
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1370
  | Equal_neg
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1371
  | Extuni_Bool2
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1372
  | Extuni_Bool1
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1373
  | Extuni_Dec
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1374
  | Extuni_Bind
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1375
  | Extuni_Triv
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1376
  | Extuni_FlexRigid
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1377
  | Extuni_Func
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1378
  | Polarity_switch
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1379
  | Forall_special_pos
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1380
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1381
datatype feature =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1382
    ConstsDiff
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1383
  | StripQuantifiers
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1384
  | Flip_Conclusion
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1385
  | Loop of loop_feature list
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1386
  | LoopOnce of loop_feature list
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1387
  | InnerLoopOnce of loop_feature list
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1388
  | CleanUp of cleanup_feature list
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1389
  | AbsorbSkolemDefs
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1390
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1391
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1392
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1393
fun can_feature x l =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1394
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1395
    fun sublist_of_clean_up el =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1396
      case el of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1397
          CleanUp l'' => SOME l''
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1398
        | _ => NONE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1399
    fun sublist_of_loop el =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1400
      case el of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1401
          Loop l'' => SOME l''
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1402
        | _ => NONE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1403
    fun sublist_of_loop_once el =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1404
      case el of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1405
          LoopOnce l'' => SOME l''
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1406
        | _ => NONE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1407
    fun sublist_of_inner_loop_once el =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1408
      case el of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1409
          InnerLoopOnce l'' => SOME l''
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1410
        | _ => NONE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1411
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1412
    fun check_sublist sought_sublist opt_list =
58412
f65f11f4854c more standard Isabelle/ML operations;
wenzelm
parents: 58411
diff changeset
  1413
      if forall is_none opt_list then false
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1414
      else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1415
        fold_options opt_list
58411
wenzelm
parents: 57773
diff changeset
  1416
        |> flat
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1417
        |> pair sought_sublist
67405
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67404
diff changeset
  1418
        |> subset (op =)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1419
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1420
    case x of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1421
        CleanUp l' =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1422
          map sublist_of_clean_up l
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1423
          |> check_sublist l'
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1424
      | Loop l' =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1425
          map sublist_of_loop l
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1426
          |> check_sublist l'
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1427
      | LoopOnce l' =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1428
          map sublist_of_loop_once l
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1429
          |> check_sublist l'
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1430
      | InnerLoopOnce l' =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1431
          map sublist_of_inner_loop_once l
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1432
          |> check_sublist l'
67405
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67404
diff changeset
  1433
      | _ => exists (curry (op =) x) l
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1434
  end;
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1435
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1436
fun loop_can_feature loop_feats l =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1437
  can_feature (Loop loop_feats) l orelse
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1438
  can_feature (LoopOnce loop_feats) l orelse
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1439
  can_feature (InnerLoopOnce loop_feats) l;
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1440
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1441
\<^assert> (can_feature ConstsDiff [StripQuantifiers, ConstsDiff]);
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1442
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1443
\<^assert>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1444
  (can_feature (CleanUp [RemoveHypothesesFromSkolemDefs])
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1445
    [CleanUp [RemoveHypothesesFromSkolemDefs, RemoveDuplicates]]);
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1446
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1447
\<^assert>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1448
  (can_feature (Loop []) [Loop [Existential_Var]]);
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1449
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1450
\<^assert>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1451
  (not (can_feature (Loop []) [InnerLoopOnce [Existential_Var]]));
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1452
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1453
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1454
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1455
exception NO_LOOP_FEATS
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1456
fun get_loop_feats (feats : feature list) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1457
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1458
    val loop_find =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1459
      fold (fn x => fn loop_feats_acc =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1460
        if is_some loop_feats_acc then loop_feats_acc
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1461
        else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1462
          case x of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1463
              Loop loop_feats => SOME loop_feats
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1464
            | LoopOnce loop_feats => SOME loop_feats
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1465
            | InnerLoopOnce loop_feats => SOME loop_feats
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1466
            | _ => NONE)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1467
       feats
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1468
       NONE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1469
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1470
    if is_some loop_find then the loop_find
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1471
    else raise NO_LOOP_FEATS
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1472
  end;
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1473
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1474
\<^assert>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1475
  (get_loop_feats [Loop [King_Cong, Break_Hypotheses, Existential_Free, Existential_Var, Universal]] =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1476
   [King_Cong, Break_Hypotheses, Existential_Free, Existential_Var, Universal])
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1477
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1478
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1479
(*use as elim rule to remove premises*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1480
lemma insa_prems: "\<lbrakk>Q; P\<rbrakk> \<Longrightarrow> P" by auto
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1481
ML \<open>
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1482
fun cleanup_skolem_defs ctxt feats =
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1483
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1484
    (*remove hypotheses from skolem defs,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1485
     after testing that they look like skolem defs*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1486
    val dehypothesise_skolem_defs =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1487
      COND' (SOME #> TERMPRED (fn _ => true) conc_is_skolem_def)
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1488
        (REPEAT_DETERM o eresolve_tac ctxt @{thms insa_prems})
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1489
        (K no_tac)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1490
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1491
    if can_feature (CleanUp [RemoveHypothesesFromSkolemDefs]) feats then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1492
      ALLGOALS (TRY o dehypothesise_skolem_defs)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1493
    else all_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1494
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1495
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1496
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1497
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1498
fun remove_duplicates_tac feats =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1499
  (if can_feature (CleanUp [RemoveDuplicates]) feats then
68825
8207c67d9ef4 clarified signature: do not expose internal operation;
wenzelm
parents: 67613
diff changeset
  1500
     distinct_subgoals_tac
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1501
   else all_tac)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1502
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1503
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1504
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1505
(*given a goal state, indicates the skolem constants committed-to in it (i.e. appearing in LHS of a skolem definition)*)
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1506
fun which_skolem_concs_used ctxt = fn st =>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1507
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1508
    val feats = [CleanUp [RemoveHypothesesFromSkolemDefs, RemoveDuplicates]]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1509
    val scrubup_tac =
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1510
      cleanup_skolem_defs ctxt feats
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1511
      THEN remove_duplicates_tac feats
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1512
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1513
    scrubup_tac st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1514
    |> break_seq
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1515
    |> tap (fn (_, rest) => \<^assert> (null (Seq.list_of rest)))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1516
    |> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1517
    |> TERMFUN (snd (*discard hypotheses*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1518
                 #> get_skolem_conc_const) NONE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1519
    |> switch (fold (fn x => fn l => if is_some x then the x :: l else l)) []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1520
    |> map Const
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1521
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1522
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1523
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1524
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1525
fun exists_tac ctxt feats consts_diff =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1526
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1527
    val ex_var =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1528
      if loop_can_feature [Existential_Var] feats andalso consts_diff <> [] then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1529
        new_skolem_tac ctxt consts_diff
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1530
        (*We're making sure that each skolem constant is used once in instantiations.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1531
      else K no_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1532
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1533
    val ex_free =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1534
      if loop_can_feature [Existential_Free] feats andalso consts_diff = [] then
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58941
diff changeset
  1535
        eresolve_tac ctxt @{thms polar_exE}
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1536
      else K no_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1537
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1538
    ex_var APPEND' ex_free
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1539
  end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1540
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1541
fun forall_tac ctxt feats =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1542
  if loop_can_feature [Universal] feats then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1543
    forall_pos_tac ctxt
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1544
  else K no_tac
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1545
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1546
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1547
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1548
subsubsection "Finite types"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1549
(*lift quantification from a singleton literal to a singleton clause*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1550
lemma forall_pos_lift:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
  1551
"\<lbrakk>(\<forall>X. P X) = True; \<forall>X. (P X = True) \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" by auto
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1552
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1553
(*predicate over the type of the leading quantified variable*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1554
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1555
ML \<open>
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1556
fun extcnf_forall_special_pos_tac ctxt =
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1557
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1558
    val bool =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1559
      ["True", "False"]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1560
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1561
    val bool_to_bool =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1562
      ["% _ . True", "% _ . False", "% x . x", "Not"]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1563
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1564
    val tacs =
59763
56d2c357e6b5 tuned signature;
wenzelm
parents: 59760
diff changeset
  1565
      map (fn t_s =>  (* FIXME proper context!? *)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1566
       Rule_Insts.eres_inst_tac \<^context> [((("x", 0), Position.none), t_s)] [] @{thm allE}
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1567
       THEN' assume_tac ctxt)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1568
  in
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1569
    (TRY o eresolve_tac ctxt @{thms forall_pos_lift})
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1570
    THEN' (assume_tac ctxt
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1571
           ORELSE' FIRST'
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1572
            (*FIXME could check the type of the leading quantified variable, instead of trying everything*)
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1573
            (tacs (bool @ bool_to_bool)))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1574
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1575
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1576
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1577
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1578
subsubsection "Emulator"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1579
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1580
lemma efq: "[|A = True; A = False|] ==> R" by auto
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1581
ML \<open>
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1582
fun efq_tac ctxt =
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1583
  (eresolve_tac ctxt @{thms efq} THEN' assume_tac ctxt)
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1584
  ORELSE' assume_tac ctxt
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1585
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1586
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1587
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1588
(*This is applied to all subgoals, repeatedly*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1589
fun extcnf_combined_main ctxt feats consts_diff =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1590
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1591
    (*This is applied to subgoals which don't have a conclusion
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1592
      consisting of a Skolem definition*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1593
    fun extcnf_combined_tac' ctxt i = fn st =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1594
      let
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1595
        val skolem_consts_used_so_far = which_skolem_concs_used ctxt st
67405
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67404
diff changeset
  1596
        val consts_diff' = subtract (op =) skolem_consts_used_so_far consts_diff
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1597
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1598
        fun feat_to_tac feat =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1599
          case feat of
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1600
              Close_Branch => trace_tac' ctxt "mark: closer" (efq_tac ctxt)
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1601
            | ConjI => trace_tac' ctxt "mark: conjI" (resolve_tac ctxt @{thms conjI})
59533
e9ffe89a20a4 proper context;
wenzelm
parents: 59498
diff changeset
  1602
            | King_Cong => trace_tac' ctxt "mark: expander_animal" (expander_animal ctxt)
e9ffe89a20a4 proper context;
wenzelm
parents: 59498
diff changeset
  1603
            | Break_Hypotheses => trace_tac' ctxt "mark: break_hypotheses" (break_hypotheses_tac ctxt)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1604
            | RemoveRedundantQuantifications => K all_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1605
(*
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1606
FIXME Building this into the loop instead.. maybe not the ideal choice
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1607
            | RemoveRedundantQuantifications =>
59533
e9ffe89a20a4 proper context;
wenzelm
parents: 59498
diff changeset
  1608
                trace_tac' ctxt "mark: strip_unused_variable_hyp"
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1609
                 (REPEAT_DETERM o remove_redundant_quantification_in_lit)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1610
*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1611
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1612
            | Assumption => assume_tac ctxt
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1613
(*FIXME both Existential_Free and Existential_Var run same code*)
59533
e9ffe89a20a4 proper context;
wenzelm
parents: 59498
diff changeset
  1614
            | Existential_Free => trace_tac' ctxt "mark: forall_neg" (exists_tac ctxt feats consts_diff')
e9ffe89a20a4 proper context;
wenzelm
parents: 59498
diff changeset
  1615
            | Existential_Var => trace_tac' ctxt "mark: forall_neg" (exists_tac ctxt feats consts_diff')
e9ffe89a20a4 proper context;
wenzelm
parents: 59498
diff changeset
  1616
            | Universal => trace_tac' ctxt "mark: forall_pos" (forall_tac ctxt feats)
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1617
            | Not_pos => trace_tac' ctxt "mark: not_pos" (dresolve_tac ctxt @{thms leo2_rules(9)})
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1618
            | Not_neg => trace_tac' ctxt "mark: not_neg" (dresolve_tac ctxt @{thms leo2_rules(10)})
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1619
            | Or_pos => trace_tac' ctxt "mark: or_pos" (dresolve_tac ctxt @{thms leo2_rules(5)}) (*could add (6) for negated conjunction*)
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1620
            | Or_neg => trace_tac' ctxt "mark: or_neg" (dresolve_tac ctxt @{thms leo2_rules(7)})
59533
e9ffe89a20a4 proper context;
wenzelm
parents: 59498
diff changeset
  1621
            | Equal_pos => trace_tac' ctxt "mark: equal_pos" (dresolve_tac ctxt (@{thms eq_pos_bool} @ [@{thm leo2_rules(3)}, @{thm eq_pos_func}]))
e9ffe89a20a4 proper context;
wenzelm
parents: 59498
diff changeset
  1622
            | Equal_neg => trace_tac' ctxt "mark: equal_neg" (dresolve_tac ctxt [@{thm eq_neg_bool}, @{thm leo2_rules(4)}])
e9ffe89a20a4 proper context;
wenzelm
parents: 59498
diff changeset
  1623
            | Donkey_Cong => trace_tac' ctxt "mark: donkey_cong" (simper_animal ctxt THEN' ex_expander_tac ctxt)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1624
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1625
            | Extuni_Bool2 => trace_tac' ctxt "mark: extuni_bool2" (dresolve_tac ctxt @{thms extuni_bool2})
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1626
            | Extuni_Bool1 => trace_tac' ctxt "mark: extuni_bool1" (dresolve_tac ctxt @{thms extuni_bool1})
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1627
            | Extuni_Bind => trace_tac' ctxt "mark: extuni_triv" (eresolve_tac ctxt @{thms extuni_triv})
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1628
            | Extuni_Triv => trace_tac' ctxt "mark: extuni_triv" (eresolve_tac ctxt @{thms extuni_triv})
59533
e9ffe89a20a4 proper context;
wenzelm
parents: 59498
diff changeset
  1629
            | Extuni_Dec => trace_tac' ctxt "mark: extuni_dec_tac" (extuni_dec_tac ctxt)
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1630
            | Extuni_FlexRigid => trace_tac' ctxt "mark: extuni_flex_rigid" (assume_tac ctxt ORELSE' asm_full_simp_tac ctxt)
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1631
            | Extuni_Func => trace_tac' ctxt "mark: extuni_func" (dresolve_tac ctxt @{thms extuni_func})
59533
e9ffe89a20a4 proper context;
wenzelm
parents: 59498
diff changeset
  1632
            | Polarity_switch => trace_tac' ctxt "mark: polarity_switch" (eresolve_tac ctxt @{thms polarity_switch})
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1633
            | Forall_special_pos => trace_tac' ctxt "mark: dorall_special_pos" (extcnf_forall_special_pos_tac ctxt)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1634
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1635
        val core_tac =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1636
          get_loop_feats feats
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1637
          |> map feat_to_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1638
          |> FIRST'
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1639
      in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1640
        core_tac i st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1641
      end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1642
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1643
    (*This is applied to all subgoals, repeatedly*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1644
    fun extcnf_combined_tac ctxt i =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1645
      COND (TERMPRED (fn _ => true) conc_is_skolem_def (SOME i))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1646
        no_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1647
        (extcnf_combined_tac' ctxt i)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1648
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1649
    val core_tac = CHANGED (ALLGOALS (IF_UNSOLVED o TRY o extcnf_combined_tac ctxt))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1650
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1651
    val full_tac = REPEAT core_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1652
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1653
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1654
    CHANGED
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1655
      (if can_feature (InnerLoopOnce []) feats then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1656
         core_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1657
       else full_tac)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1658
  end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1659
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1660
val interpreted_consts =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1661
  [\<^const_name>\<open>HOL.All\<close>, \<^const_name>\<open>HOL.Ex\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1662
   \<^const_name>\<open>Hilbert_Choice.Eps\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1663
   \<^const_name>\<open>HOL.conj\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1664
   \<^const_name>\<open>HOL.disj\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1665
   \<^const_name>\<open>HOL.eq\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1666
   \<^const_name>\<open>HOL.implies\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1667
   \<^const_name>\<open>HOL.The\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1668
   \<^const_name>\<open>HOL.Ex1\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1669
   \<^const_name>\<open>HOL.Not\<close>,
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1670
   (* @{const_name HOL.iff}, *) (*FIXME do these exist?*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1671
   (* @{const_name HOL.not_equal}, *)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1672
   \<^const_name>\<open>HOL.False\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1673
   \<^const_name>\<open>HOL.True\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1674
   \<^const_name>\<open>Pure.imp\<close>]
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1675
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1676
fun strip_qtfrs_tac ctxt =
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1677
  REPEAT_DETERM (HEADGOAL (resolve_tac ctxt @{thms allI}))
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1678
  THEN REPEAT_DETERM (HEADGOAL (eresolve_tac ctxt @{thms exE}))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1679
  THEN HEADGOAL (canonicalise_qtfr_order ctxt)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1680
  THEN
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1681
    ((REPEAT (HEADGOAL (nominal_inst_parametermatch_tac ctxt @{thm allE})))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1682
     APPEND (REPEAT (HEADGOAL (inst_parametermatch_tac ctxt [@{thm allE}]))))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1683
  (*FIXME need to handle "@{thm exI}"?*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1684
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1685
(*difference in constants between the hypothesis clause and the conclusion clause*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1686
fun clause_consts_diff thm =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1687
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1688
    val t =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59533
diff changeset
  1689
      Thm.prop_of thm
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1690
      |> Logic.dest_implies
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1691
      |> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1692
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1693
      (*This bit should not be needed, since Leo2 inferences don't have parameters*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1694
      |> TPTP_Reconstruct.strip_top_all_vars []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1695
      |> snd
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1696
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1697
    val do_diff =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1698
      Logic.dest_implies
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1699
      #> uncurry TPTP_Reconstruct.new_consts_between
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1700
      #> filter
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1701
           (fn Const (n, _) =>
67405
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67404
diff changeset
  1702
             not (member (op =) interpreted_consts n))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1703
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1704
    if head_of t = Logic.implies then do_diff t
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1705
    else []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1706
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1707
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1708
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1709
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1710
(*remove quantification in hypothesis clause (! X. t), if
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1711
  X not free in t*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1712
fun remove_redundant_quantification ctxt i = fn st =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1713
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1714
    val gls =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59533
diff changeset
  1715
      Thm.prop_of st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1716
      |> Logic.strip_horn
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1717
      |> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1718
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1719
    if null gls then raise NO_GOALS
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1720
    else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1721
      let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1722
        val (params, (hyp_clauses, conc_clause)) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1723
          rpair (i - 1) gls
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1724
          |> uncurry nth
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1725
          |> TPTP_Reconstruct.strip_top_all_vars []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1726
          |> apsnd Logic.strip_horn
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1727
      in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1728
        (*this is to fail gracefully in case this tactic is applied to a goal which doesn't have a single hypothesis*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1729
        if length hyp_clauses > 1 then no_tac st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1730
        else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1731
          let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1732
            val hyp_clause = the_single hyp_clauses
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1733
            val sep_prefix =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1734
              HOLogic.dest_Trueprop
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1735
              #> TPTP_Reconstruct.strip_top_All_vars
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1736
              #> apfst rev
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1737
            val (hyp_prefix, hyp_body) = sep_prefix hyp_clause
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1738
            val (conc_prefix, conc_body) = sep_prefix conc_clause
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1739
          in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1740
            if null hyp_prefix orelse
67405
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67404
diff changeset
  1741
              member (op =) conc_prefix (hd hyp_prefix) orelse
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67404
diff changeset
  1742
              member (op =)  (Term.add_frees hyp_body []) (hd hyp_prefix) then
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1743
              no_tac st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1744
            else
59780
23b67731f4f0 support 'for' fixes in rule_tac etc.;
wenzelm
parents: 59763
diff changeset
  1745
              Rule_Insts.eres_inst_tac ctxt [((("x", 0), Position.none), "(@X. False)")] []
59755
f8d164ab0dc1 more position information;
wenzelm
parents: 59639
diff changeset
  1746
                @{thm allE} i st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1747
          end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1748
     end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1749
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1750
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1751
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1752
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1753
fun remove_redundant_quantification_ignore_skolems ctxt i =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1754
  COND (TERMPRED (fn _ => true) conc_is_skolem_def (SOME i))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1755
    no_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1756
    (remove_redundant_quantification ctxt i)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1757
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1758
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1759
lemma drop_redundant_literal_qtfr:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
  1760
  "(\<forall>X. P) = True \<Longrightarrow> P = True"
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
  1761
  "(\<exists>X. P) = True \<Longrightarrow> P = True"
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
  1762
  "(\<forall>X. P) = False \<Longrightarrow> P = False"
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
  1763
  "(\<exists>X. P) = False \<Longrightarrow> P = False"
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1764
by auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1765
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1766
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1767
(*remove quantification in the literal "(! X. t) = True/False"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1768
  in the singleton hypothesis clause, if X not free in t*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1769
fun remove_redundant_quantification_in_lit ctxt i = fn st =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1770
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1771
    val gls =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59533
diff changeset
  1772
      Thm.prop_of st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1773
      |> Logic.strip_horn
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1774
      |> fst
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1775
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1776
    if null gls then raise NO_GOALS
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1777
    else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1778
      let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1779
        val (params, (hyp_clauses, conc_clause)) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1780
          rpair (i - 1) gls
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1781
          |> uncurry nth
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1782
          |> TPTP_Reconstruct.strip_top_all_vars []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1783
          |> apsnd Logic.strip_horn
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1784
      in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1785
        (*this is to fail gracefully in case this tactic is applied to a goal which doesn't have a single hypothesis*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1786
        if length hyp_clauses > 1 then no_tac st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1787
        else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1788
          let
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1789
            fun literal_content (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ lhs $ (rhs as \<^term>\<open>True\<close>)) = SOME (lhs, rhs)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  1790
              | literal_content (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ lhs $ (rhs as \<^term>\<open>False\<close>)) = SOME (lhs, rhs)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1791
              | literal_content t = NONE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1792
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1793
            val hyp_clause =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1794
              the_single hyp_clauses
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1795
              |> HOLogic.dest_Trueprop
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1796
              |> literal_content
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1797
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1798
          in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1799
            if is_none hyp_clause then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1800
              no_tac st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1801
            else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1802
              let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1803
                val (hyp_lit_prefix, hyp_lit_body) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1804
                  the hyp_clause
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1805
                  |> (fn (t, polarity) =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1806
                       TPTP_Reconstruct.strip_top_All_vars t
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1807
                       |> apfst rev)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1808
              in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1809
                if null hyp_lit_prefix orelse
67405
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67404
diff changeset
  1810
                  member (op =)  (Term.add_frees hyp_lit_body []) (hd hyp_lit_prefix) then
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1811
                  no_tac st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1812
                else
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58941
diff changeset
  1813
                  dresolve_tac ctxt @{thms drop_redundant_literal_qtfr} i st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1814
              end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1815
          end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1816
     end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1817
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1818
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1819
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1820
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1821
fun remove_redundant_quantification_in_lit_ignore_skolems ctxt i =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1822
  COND (TERMPRED (fn _ => true) conc_is_skolem_def (SOME i))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1823
    no_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1824
    (remove_redundant_quantification_in_lit ctxt i)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1825
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1826
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1827
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1828
fun extcnf_combined_tac ctxt prob_name_opt feats skolem_consts = fn st =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1829
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1830
    val thy = Proof_Context.theory_of ctxt
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1831
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1832
    (*Initially, st consists of a single goal, showing the
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1833
      hypothesis clause implying the conclusion clause.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1834
      There are no parameters.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1835
    val consts_diff =
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 63167
diff changeset
  1836
      union (=) skolem_consts
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1837
       (if can_feature ConstsDiff feats then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1838
          clause_consts_diff st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1839
        else [])
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1840
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1841
    val main_tac =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1842
      if can_feature (LoopOnce []) feats orelse can_feature (InnerLoopOnce []) feats then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1843
        extcnf_combined_main ctxt feats consts_diff
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1844
      else if can_feature (Loop []) feats then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1845
        BEST_FIRST (TERMPRED (fn _ => true) conc_is_skolem_def NONE, size_of_thm)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1846
(*FIXME maybe need to weaken predicate to include "solved form"?*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1847
         (extcnf_combined_main ctxt feats consts_diff)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1848
      else all_tac (*to allow us to use the cleaning features*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1849
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1850
    (*Remove hypotheses from Skolem definitions,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1851
      then remove duplicate subgoals,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1852
      then we should be left with skolem definitions:
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1853
        absorb them as axioms into the theory.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1854
    val cleanup =
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1855
      cleanup_skolem_defs ctxt feats
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1856
      THEN remove_duplicates_tac feats
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1857
      THEN (if can_feature AbsorbSkolemDefs feats then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1858
              ALLGOALS (absorb_skolem_def ctxt prob_name_opt)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1859
            else all_tac)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1860
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1861
    val have_loop_feats =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1862
      (get_loop_feats feats; true)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1863
      handle NO_LOOP_FEATS => false
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1864
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1865
    val tec =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1866
      (if can_feature StripQuantifiers feats then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1867
         (REPEAT (CHANGED (strip_qtfrs_tac ctxt)))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1868
       else all_tac)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1869
      THEN (if can_feature Flip_Conclusion feats then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1870
             HEADGOAL (flip_conclusion_tac ctxt)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1871
           else all_tac)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1872
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1873
      (*after stripping the quantifiers any remaining quantifiers
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1874
        can be simply eliminated -- they're redundant*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1875
      (*FIXME instead of just using allE, instantiate to a silly
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1876
         term, to remove opportunities for unification.*)
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1877
      THEN (REPEAT_DETERM (eresolve_tac ctxt @{thms allE} 1))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1878
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1879
      THEN (REPEAT_DETERM (resolve_tac ctxt @{thms allI} 1))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1880
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1881
      THEN (if have_loop_feats then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1882
              REPEAT (CHANGED
59533
e9ffe89a20a4 proper context;
wenzelm
parents: 59498
diff changeset
  1883
              ((ALLGOALS (TRY o clause_breaker_tac ctxt)) (*brush away literals which don't change*)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1884
               THEN
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1885
                (*FIXME move this to a different level?*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1886
                (if loop_can_feature [Polarity_switch] feats then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1887
                   all_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1888
                 else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1889
                   (TRY (IF_UNSOLVED (HEADGOAL (remove_redundant_quantification_ignore_skolems ctxt))))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1890
                   THEN (TRY (IF_UNSOLVED (HEADGOAL (remove_redundant_quantification_in_lit_ignore_skolems ctxt)))))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1891
               THEN (TRY main_tac)))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1892
            else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1893
              all_tac)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1894
      THEN IF_UNSOLVED cleanup
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1895
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1896
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1897
    DEPTH_SOLVE (CHANGED tec) st
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1898
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1899
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1900
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1901
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1902
subsubsection "unfold_def"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1903
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1904
(*this is used when handling unfold_tac, because the skeleton includes the definitions conjoined with the goal. it turns out that, for my tactic, the definitions are harmful. instead of modifying the skeleton (which may be nontrivial) i'm just dropping the information using this lemma. obviously, and from the name, order matters here.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1905
lemma drop_first_hypothesis [rule_format]: "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> B" by auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1906
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1907
(*Unfold_def works by reducing the goal to a meta equation,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1908
  then working on it until it can be discharged by atac,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1909
  or reflexive, or else turned back into an object equation
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1910
  and broken down further.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1911
lemma un_meta_polarise: "(X \<equiv> True) \<Longrightarrow> X" by auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1912
lemma meta_polarise: "X \<Longrightarrow> X \<equiv> True" by auto
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1913
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1914
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1915
fun unfold_def_tac ctxt depends_on_defs = fn st =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1916
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1917
    (*This is used when we end up with something like
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1918
        (A & B) \<equiv> True \<Longrightarrow> (B & A) \<equiv> True.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1919
      It breaks down this subgoal until it can be trivially
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1920
      discharged.
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1921
     *)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1922
    val kill_meta_eqs_tac =
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1923
      dresolve_tac ctxt @{thms un_meta_polarise}
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1924
      THEN' resolve_tac ctxt @{thms meta_polarise}
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1925
      THEN' (REPEAT_DETERM o (eresolve_tac ctxt @{thms conjE}))
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1926
      THEN' (REPEAT_DETERM o (resolve_tac ctxt @{thms conjI} ORELSE' assume_tac ctxt))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1927
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1928
    val continue_reducing_tac =
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1929
      resolve_tac ctxt @{thms meta_eq_to_obj_eq} 1
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1930
      THEN (REPEAT_DETERM (ex_expander_tac ctxt 1))
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1931
      THEN TRY (polarise_subgoal_hyps ctxt 1) (*no need to REPEAT_DETERM here, since there should only be one hypothesis*)
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1932
      THEN TRY (dresolve_tac ctxt @{thms eq_reflection} 1)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1933
      THEN (TRY ((CHANGED o rewrite_goal_tac ctxt
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1934
              (@{thm expand_iff} :: @{thms simp_meta})) 1))
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1935
      THEN HEADGOAL (resolve_tac ctxt @{thms reflexive}
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1936
                     ORELSE' assume_tac ctxt
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1937
                     ORELSE' kill_meta_eqs_tac)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1938
60318
ab785001b51d tuned spelling;
wenzelm
parents: 59882
diff changeset
  1939
    val tactic =
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1940
      (resolve_tac ctxt @{thms polarise} 1 THEN assume_tac ctxt 1)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1941
      ORELSE
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1942
        (REPEAT_DETERM (eresolve_tac ctxt @{thms conjE} 1 THEN
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1943
          eresolve_tac ctxt @{thms drop_first_hypothesis} 1)
60319
127c2f00ca94 standardize towards Thm.eta_long_conversion, which just does eta_long conversion;
wenzelm
parents: 60318
diff changeset
  1944
         THEN PRIMITIVE (Conv.fconv_rule Thm.eta_long_conversion)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1945
         THEN (REPEAT_DETERM (ex_expander_tac ctxt 1))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1946
         THEN (TRY ((CHANGED o rewrite_goal_tac ctxt @{thms simp_meta}) 1))
60319
127c2f00ca94 standardize towards Thm.eta_long_conversion, which just does eta_long conversion;
wenzelm
parents: 60318
diff changeset
  1947
         THEN PRIMITIVE (Conv.fconv_rule Thm.eta_long_conversion)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1948
         THEN
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  1949
           (HEADGOAL (assume_tac ctxt)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1950
           ORELSE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1951
            (unfold_tac ctxt depends_on_defs
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1952
             THEN IF_UNSOLVED continue_reducing_tac)))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1953
  in
60318
ab785001b51d tuned spelling;
wenzelm
parents: 59882
diff changeset
  1954
    tactic st
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1955
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1956
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1957
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1958
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1959
subsection "Handling split 'preprocessing'"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1960
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1961
lemma split_tranfs:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
  1962
  "\<forall>x. P x \<and> Q x \<equiv> (\<forall>x. P x) \<and> (\<forall>x. Q x)"
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
  1963
  "\<not> (\<not> A) \<equiv> A"
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
  1964
  "\<exists>x. A \<equiv> A"
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
  1965
  "(A \<and> B) \<and> C \<equiv> A \<and> B \<and> C"
ce654b0e6d69 more symbols;
wenzelm
parents: 67405
diff changeset
  1966
  "A = B \<equiv> (A \<longrightarrow> B) \<and> (B \<longrightarrow> A)"
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1967
by (rule eq_reflection, auto)+
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1968
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1969
(*Same idiom as ex_expander_tac*)
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1970
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1971
fun split_simp_tac (ctxt : Proof.context) i =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1972
   let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1973
     val simpset =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1974
       fold Simplifier.add_simp @{thms split_tranfs} (empty_simpset ctxt)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1975
   in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1976
     CHANGED (asm_full_simp_tac simpset i)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1977
   end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1978
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1979
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1980
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1981
subsection "Alternative reconstruction tactics"
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1982
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1983
(*An "auto"-based proof reconstruction, where we attempt to reconstruct each inference
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1984
  using auto_tac. A realistic tactic would inspect the inference name and act
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1985
  accordingly.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1986
fun auto_based_reconstruction_tac ctxt prob_name n =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1987
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1988
    val thy = Proof_Context.theory_of ctxt
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1989
    val pannot = TPTP_Reconstruct.get_pannot_of_prob thy prob_name
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1990
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1991
    TPTP_Reconstruct.inference_at_node
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1992
     thy
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1993
     prob_name (#meta pannot) n
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1994
      |> the
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1995
      |> (fn {inference_fmla, ...} =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1996
          Goal.prove ctxt [] [] inference_fmla
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1997
           (fn pdata => auto_tac (#context pdata)))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  1998
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  1999
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2000
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2001
(*An oracle-based reconstruction, which is only used to test the shunting part of the system*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2002
oracle oracle_iinterp = "fn t => t"
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  2003
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2004
fun oracle_based_reconstruction_tac ctxt prob_name n =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2005
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2006
    val thy = Proof_Context.theory_of ctxt
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2007
    val pannot = TPTP_Reconstruct.get_pannot_of_prob thy prob_name
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2008
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2009
    TPTP_Reconstruct.inference_at_node
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2010
     thy
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2011
     prob_name (#meta pannot) n
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2012
      |> the
59639
f596ed647018 clarified context;
wenzelm
parents: 59621
diff changeset
  2013
      |> (fn {inference_fmla, ...} => Thm.cterm_of ctxt inference_fmla)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2014
      |> oracle_iinterp
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2015
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  2016
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2017
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2018
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2019
subsection "Leo2 reconstruction tactic"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2020
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  2021
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2022
(*Failure reports can be adjusted to avoid interrupting
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2023
  an overall reconstruction process*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2024
fun fail ctxt x =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2025
  if unexceptional_reconstruction ctxt then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2026
    (warning x; raise INTERPRET_INFERENCE)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2027
  else error x
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2028
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2029
fun interpret_leo2_inference_tac ctxt prob_name node =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2030
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2031
    val thy = Proof_Context.theory_of ctxt
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2032
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2033
    val _ =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2034
      if Config.get ctxt tptp_trace_reconstruction then
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2035
        tracing ("interpret_inference reconstructing node" ^ node ^ " of " ^ TPTP_Problem_Name.mangle_problem_name prob_name)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2036
      else ()
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2037
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2038
    val pannot = TPTP_Reconstruct.get_pannot_of_prob thy prob_name
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2039
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2040
    fun nonfull_extcnf_combined_tac feats =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2041
      extcnf_combined_tac ctxt (SOME prob_name)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2042
       [ConstsDiff,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2043
        StripQuantifiers,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2044
        InnerLoopOnce (Break_Hypotheses :: (*FIXME RemoveRedundantQuantifications :: *) feats),
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2045
        AbsorbSkolemDefs]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2046
       []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2047
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2048
    val source_inf_opt =
67405
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67404
diff changeset
  2049
      AList.lookup (op =) (#meta pannot)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2050
      #> the
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2051
      #> #source_inf_opt
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2052
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2053
    (*FIXME integrate this with other lookup code, or in the early analysis*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2054
    local
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2055
      fun node_is_of_role role node =
67405
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67404
diff changeset
  2056
        AList.lookup (op =) (#meta pannot) node |> the
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2057
        |> #role
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2058
        |> (fn role' => role = role')
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2059
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2060
      fun roled_dependencies_names role =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2061
        let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2062
          fun values () =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2063
            case role of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2064
                TPTP_Syntax.Role_Definition =>
59858
890b68e1e8b6 support for strictly private name space entries;
wenzelm
parents: 59780
diff changeset
  2065
                  map (apsnd Binding.name_of) (#defs pannot)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2066
              | TPTP_Syntax.Role_Axiom =>
59858
890b68e1e8b6 support for strictly private name space entries;
wenzelm
parents: 59780
diff changeset
  2067
                  map (apsnd Binding.name_of) (#axs pannot)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2068
              | _ => raise UNSUPPORTED_ROLE
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2069
          in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2070
            if is_none (source_inf_opt node) then []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2071
            else
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2072
              case the (source_inf_opt node) of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2073
                  TPTP_Proof.Inference (_, _, parent_inf) =>
58412
f65f11f4854c more standard Isabelle/ML operations;
wenzelm
parents: 58411
diff changeset
  2074
                    map TPTP_Proof.parent_name parent_inf
f65f11f4854c more standard Isabelle/ML operations;
wenzelm
parents: 58411
diff changeset
  2075
                    |> filter (node_is_of_role role)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2076
                    |> (*FIXME currently definitions are not
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2077
                         included in the proof annotations, so
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2078
                         i'm using all the definitions available
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2079
                         in the proof. ideally i should only
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2080
                         use the ones in the proof annotation.*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2081
                       (fn x =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2082
                         if role = TPTP_Syntax.Role_Definition then
59858
890b68e1e8b6 support for strictly private name space entries;
wenzelm
parents: 59780
diff changeset
  2083
                           let fun values () = map (apsnd Binding.name_of) (#defs pannot)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2084
                           in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2085
                             map snd (values ())
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2086
                           end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2087
                         else
67405
e9ab4ad7bd15 uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents: 67404
diff changeset
  2088
                         map (fn node => AList.lookup (op =) (values ()) node |> the) x)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2089
                | _ => []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2090
         end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2091
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2092
      val roled_dependencies =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2093
        roled_dependencies_names
59858
890b68e1e8b6 support for strictly private name space entries;
wenzelm
parents: 59780
diff changeset
  2094
        #> map (Global_Theory.get_thm thy)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2095
    in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2096
      val depends_on_defs = roled_dependencies TPTP_Syntax.Role_Definition
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2097
      val depends_on_axs = roled_dependencies TPTP_Syntax.Role_Axiom
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2098
      val depends_on_defs_names = roled_dependencies_names TPTP_Syntax.Role_Definition
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2099
    end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2100
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2101
    fun get_binds source_inf_opt =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2102
      case the source_inf_opt of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2103
          TPTP_Proof.Inference (_, _, parent_inf) =>
58411
wenzelm
parents: 57773
diff changeset
  2104
            maps
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2105
              (fn TPTP_Proof.Parent _ => []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2106
                | TPTP_Proof.ParentWithDetails (_, parent_details) => parent_details)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2107
              parent_inf
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2108
        | _ => []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2109
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2110
    val inference_name =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2111
      case TPTP_Reconstruct.inference_at_node thy prob_name (#meta pannot) node of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2112
          NONE => fail ctxt "Cannot reconstruct rule: no information"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2113
        | SOME {inference_name, ...} => inference_name
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2114
    val default_tac = HEADGOAL (blast_tac ctxt)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2115
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2116
    case inference_name of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2117
      "fo_atp_e" =>
57773
2719eb9d40fe Metis is being used to emulate E steps;
sultana
parents: 56252
diff changeset
  2118
        HEADGOAL (Metis_Tactic.metis_tac [] ATP_Problem_Generate.combs_or_liftingN ctxt [])
2719eb9d40fe Metis is being used to emulate E steps;
sultana
parents: 56252
diff changeset
  2119
        (*NOTE To treat E as an oracle use the following line:
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2120
        HEADGOAL (etac (oracle_based_reconstruction_tac ctxt prob_name node))
57773
2719eb9d40fe Metis is being used to emulate E steps;
sultana
parents: 56252
diff changeset
  2121
        *)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2122
    | "copy" =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2123
         HEADGOAL
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  2124
          (assume_tac ctxt
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2125
           ORELSE'
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  2126
              (resolve_tac ctxt @{thms polarise}
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  2127
               THEN' assume_tac ctxt))
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2128
    | "polarity_switch" => nonfull_extcnf_combined_tac [Polarity_switch]
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  2129
    | "solved_all_splits" => solved_all_splits_tac ctxt
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2130
    | "extcnf_not_pos" => nonfull_extcnf_combined_tac [Not_pos]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2131
    | "extcnf_forall_pos" => nonfull_extcnf_combined_tac [Universal]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2132
    | "negate_conjecture" => fail ctxt "Should not handle negate_conjecture here"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2133
    | "unfold_def" => unfold_def_tac ctxt depends_on_defs
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2134
    | "extcnf_not_neg" => nonfull_extcnf_combined_tac [Not_neg]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2135
    | "extcnf_or_neg" => nonfull_extcnf_combined_tac [Or_neg]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2136
    | "extcnf_equal_pos" => nonfull_extcnf_combined_tac [Equal_pos]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2137
    | "extcnf_equal_neg" => nonfull_extcnf_combined_tac [Equal_neg]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2138
    | "extcnf_forall_special_pos" =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2139
         nonfull_extcnf_combined_tac [Forall_special_pos]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2140
         ORELSE HEADGOAL (blast_tac ctxt)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2141
    | "extcnf_or_pos" => nonfull_extcnf_combined_tac [Or_pos]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2142
    | "extuni_bool2" => nonfull_extcnf_combined_tac [Extuni_Bool2]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2143
    | "extuni_bool1" => nonfull_extcnf_combined_tac [Extuni_Bool1]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2144
    | "extuni_dec" =>
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  2145
        HEADGOAL (assume_tac ctxt)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2146
        ORELSE nonfull_extcnf_combined_tac [Extuni_Dec]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2147
    | "extuni_bind" => nonfull_extcnf_combined_tac [Extuni_Bind]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2148
    | "extuni_triv" => nonfull_extcnf_combined_tac [Extuni_Triv]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2149
    | "extuni_flex_rigid" => nonfull_extcnf_combined_tac [Extuni_FlexRigid]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2150
    | "prim_subst" => nonfull_extcnf_combined_tac [Assumption]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2151
    | "bind" =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2152
        let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2153
          val ordered_binds = get_binds (source_inf_opt node)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2154
        in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2155
          bind_tac ctxt prob_name ordered_binds
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2156
        end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2157
    | "standard_cnf" => HEADGOAL (standard_cnf_tac ctxt)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2158
    | "extcnf_forall_neg" =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2159
        nonfull_extcnf_combined_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2160
         [Existential_Var(* , RemoveRedundantQuantifications *)] (*FIXME RemoveRedundantQuantifications*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2161
    | "extuni_func" =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2162
        nonfull_extcnf_combined_tac [Extuni_Func, Existential_Var]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2163
    | "replace_leibnizEQ" => nonfull_extcnf_combined_tac [Assumption]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2164
    | "replace_andrewsEQ" => nonfull_extcnf_combined_tac [Assumption]
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2165
    | "split_preprocessing" =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2166
         (REPEAT (HEADGOAL (split_simp_tac ctxt)))
60319
127c2f00ca94 standardize towards Thm.eta_long_conversion, which just does eta_long conversion;
wenzelm
parents: 60318
diff changeset
  2167
         THEN TRY (PRIMITIVE (Conv.fconv_rule Thm.eta_long_conversion))
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
  2168
         THEN HEADGOAL (assume_tac ctxt)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2169
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2170
    (*FIXME some of these could eventually be handled specially*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2171
    | "fac_restr" => default_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2172
    | "sim" => default_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2173
    | "res" => default_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2174
    | "rename" => default_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2175
    | "flexflex" => default_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2176
    | other => fail ctxt ("Unknown inference rule: " ^ other)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2177
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  2178
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2179
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  2180
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2181
fun interpret_leo2_inference ctxt prob_name node =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2182
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2183
    val thy = Proof_Context.theory_of ctxt
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2184
    val pannot = TPTP_Reconstruct.get_pannot_of_prob thy prob_name
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2185
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2186
    val (inference_name, inference_fmla) =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2187
      case TPTP_Reconstruct.inference_at_node thy prob_name (#meta pannot) node of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2188
          NONE => fail ctxt "Cannot reconstruct rule: no information"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2189
        | SOME {inference_name, inference_fmla, ...} =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2190
            (inference_name, inference_fmla)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2191
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2192
    val proof_outcome =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2193
      let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2194
        fun prove () =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2195
          Goal.prove ctxt [] [] inference_fmla
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2196
           (fn pdata => interpret_leo2_inference_tac
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2197
            (#context pdata) prob_name node)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2198
      in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2199
        if informative_failure ctxt then SOME (prove ())
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2200
        else try prove ()
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2201
      end
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2202
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2203
  in case proof_outcome of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2204
      NONE => fail ctxt (Pretty.string_of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2205
        (Pretty.block
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2206
          [Pretty.str ("Failed inference reconstruction for '" ^
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2207
            inference_name ^ "' at node " ^ node ^ ":\n"),
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2208
           Syntax.pretty_term ctxt inference_fmla]))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2209
    | SOME thm => thm
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2210
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  2211
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2212
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  2213
ML \<open>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2214
(*filter a set of nodes based on which inference rule was used to
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2215
  derive a node*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2216
fun nodes_by_inference (fms : TPTP_Reconstruct.formula_meaning list) inference_rule =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2217
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2218
    fun fold_fun n l =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2219
      case TPTP_Reconstruct.node_info fms #source_inf_opt n of
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2220
          NONE => l
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2221
        | SOME (TPTP_Proof.File _) => l
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2222
        | SOME (TPTP_Proof.Inference (rule_name, _, _)) =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2223
            if rule_name = inference_rule then n :: l
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2224
            else l
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2225
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2226
    fold fold_fun (map fst fms) []
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2227
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  2228
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2229
55597
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2230
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2231
section "Importing proofs and reconstructing theorems"
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2232
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  2233
ML \<open>
55597
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2234
(*Preprocessing carried out on a LEO-II proof.*)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2235
fun leo2_on_load (pannot : TPTP_Reconstruct.proof_annotation) thy =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2236
  let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2237
    val ctxt = Proof_Context.init_global thy
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68825
diff changeset
  2238
    val dud = ("", Binding.empty, \<^term>\<open>False\<close>)
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2239
    val pre_skolem_defs =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2240
      nodes_by_inference (#meta pannot) "extcnf_forall_neg" @
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2241
       nodes_by_inference (#meta pannot) "extuni_func"
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2242
      |> map (fn x =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2243
              (interpret_leo2_inference ctxt (#problem_name pannot) x; dud)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2244
               handle NO_SKOLEM_DEF (s, bnd, t) => (s, bnd, t))
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2245
      |> filter (fn (x, _, _) => x <> "") (*In case no skolem constants were introduced in that inference*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2246
    val skolem_defs = map (fn (x, y, _) => (x, y)) pre_skolem_defs
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2247
    val thy' =
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2248
      fold (fn skolem_def => fn thy =>
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2249
             let
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2250
               val ((s, thm), thy') = Thm.add_axiom_global skolem_def thy
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2251
               (* val _ = warning ("Added skolem definition " ^ s ^ ": " ^  @{make_string thm}) *) (*FIXME use of make_string*)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2252
             in thy' end)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2253
       (map (fn (_, y, z) => (y, z)) pre_skolem_defs)
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2254
       thy
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2255
  in
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2256
    ({problem_name = #problem_name pannot,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2257
      skolem_defs = skolem_defs,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2258
      defs = #defs pannot,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2259
      axs = #axs pannot,
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2260
      meta = #meta pannot},
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2261
     thy')
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2262
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  2263
\<close>
55596
928b9f677165 reconstruction framework for LEO-II's TPTP proofs;
sultana
parents:
diff changeset
  2264
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  2265
ML \<open>
55597
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2266
(*Imports and reconstructs a LEO-II proof.*)
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2267
fun reconstruct_leo2 path thy =
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2268
  let
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2269
    val prob_file = Path.base path
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2270
    val dir = Path.dir path
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2271
    val thy' = TPTP_Reconstruct.import_thm true [dir, prob_file] path leo2_on_load thy
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2272
    val ctxt =
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2273
      Context.Theory thy'
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2274
      |> Context.proof_of
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2275
    val prob_name =
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2276
      Path.implode prob_file
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2277
      |> TPTP_Problem_Name.parse_problem_name
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2278
    val theorem =
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2279
      TPTP_Reconstruct.reconstruct ctxt
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2280
       (TPTP_Reconstruct.naive_reconstruct_tac ctxt interpret_leo2_inference)
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2281
       prob_name
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2282
  in
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2283
    (*NOTE we could return the theorem value alone, since
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2284
       users could get the thy value from the thm value.*)
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2285
    (thy', theorem)
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2286
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62243
diff changeset
  2287
\<close>
55597
25d7b485df81 added a function that carries out all the reconstruction steps, for improved usability;
sultana
parents: 55596
diff changeset
  2288
57773
2719eb9d40fe Metis is being used to emulate E steps;
sultana
parents: 56252
diff changeset
  2289
end