src/HOL/Tools/res_atp.ML
author paulson
Tue, 21 Jun 2005 13:34:24 +0200
changeset 16515 7896ea4f3a87
parent 16478 d0a1f6231e2f
child 16520 7a9cda53bfa2
permissions -rw-r--r--
VAMPIRE_HOME, helper_path and various stylistic tweaks
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
     1
(*  Author: Jia Meng, Cambridge University Computer Laboratory
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
     2
    ID: $Id$
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
     3
    Copyright 2004 University of Cambridge
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     4
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     5
ATPs with TPTP format input.
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     6
*)
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
     7
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
     8
(*Jia: changed: isar_atp now processes entire proof context.  fetch thms from delta simpset/claset*)
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
     9
(*Claire: changed: added actual watcher calls *)
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
    10
15779
aed221aff642 Removed remaining references to Main.thy in reconstruction code.
quigley
parents: 15774
diff changeset
    11
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    12
signature RES_ATP = 
16478
d0a1f6231e2f Added VampCommunication.ML.
quigley
parents: 16475
diff changeset
    13
sig  
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    14
val trace_res : bool ref
15779
aed221aff642 Removed remaining references to Main.thy in reconstruction code.
quigley
parents: 15774
diff changeset
    15
val subgoals: Thm.thm list
aed221aff642 Removed remaining references to Main.thy in reconstruction code.
quigley
parents: 15774
diff changeset
    16
val traceflag : bool ref
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    17
val axiom_file : Path.T
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    18
val hyps_file : Path.T
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    19
val isar_atp : ProofContext.context * Thm.thm -> unit
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
    20
val prob_file : Path.T;
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    21
(*val atp_ax_tac : Thm.thm list -> int -> Tactical.tactic*)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    22
(*val atp_tac : int -> Tactical.tactic*)
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    23
val debug: bool ref
16156
2f6fc19aba1e Fixed array containing clasimpset rules. Added flags to turn on and off reconstruction and full spass
quigley
parents: 16089
diff changeset
    24
val full_spass: bool ref
16478
d0a1f6231e2f Added VampCommunication.ML.
quigley
parents: 16475
diff changeset
    25
(*val spass: bool ref*)
d0a1f6231e2f Added VampCommunication.ML.
quigley
parents: 16475
diff changeset
    26
val vampire: bool ref
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    27
end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    28
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    29
structure ResAtp : RES_ATP =
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    30
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    31
struct
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    32
15779
aed221aff642 Removed remaining references to Main.thy in reconstruction code.
quigley
parents: 15774
diff changeset
    33
val subgoals = [];
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    34
15779
aed221aff642 Removed remaining references to Main.thy in reconstruction code.
quigley
parents: 15774
diff changeset
    35
val traceflag = ref true;
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
    36
(* used for debug *)
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
    37
val debug = ref false;
16172
629ba53a072f small tweaks; also now write_out_clasimp takes the current theory as argument
paulson
parents: 16156
diff changeset
    38
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    39
fun debug_tac tac = (warning "testing";tac);
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
    40
(* default value is false *)
16478
d0a1f6231e2f Added VampCommunication.ML.
quigley
parents: 16475
diff changeset
    41
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
    42
val full_spass = ref false;
16478
d0a1f6231e2f Added VampCommunication.ML.
quigley
parents: 16475
diff changeset
    43
d0a1f6231e2f Added VampCommunication.ML.
quigley
parents: 16475
diff changeset
    44
(* use spass as default prover *)
d0a1f6231e2f Added VampCommunication.ML.
quigley
parents: 16475
diff changeset
    45
(*val spass = ref true;*)
d0a1f6231e2f Added VampCommunication.ML.
quigley
parents: 16475
diff changeset
    46
d0a1f6231e2f Added VampCommunication.ML.
quigley
parents: 16475
diff changeset
    47
val vampire = ref false;
d0a1f6231e2f Added VampCommunication.ML.
quigley
parents: 16475
diff changeset
    48
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    49
val trace_res = ref false;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    50
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    51
val skolem_tac = skolemize_tac;
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    52
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    53
val num_of_clauses = ref 0;
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    54
val clause_arr = Array.array(3500, ("empty", 0));
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    55
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    56
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    57
val atomize_tac =
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    58
    SUBGOAL
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    59
     (fn (prop,_) =>
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    60
	 let val ts = Logic.strip_assums_hyp prop
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    61
	 in EVERY1 
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    62
		[METAHYPS
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    63
		     (fn hyps => (cut_facts_tac (map (ObjectLogic.atomize_thm o forall_intr_vars) hyps) 1)),
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    64
	  REPEAT_DETERM_N (length ts) o (etac thin_rl)]
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    65
     end);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    66
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    67
(* temporarily use these files, which will be loaded by Vampire *)
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    68
val file_id_num =ref 0;
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    69
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    70
fun new_prob_file () =  (file_id_num := (!file_id_num) + 1;"prob"^(string_of_int (!file_id_num)));
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    71
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    72
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    73
val axiom_file = File.tmp_path (Path.basic "axioms");
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    74
val clasimp_file = File.tmp_path (Path.basic "clasimp");
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    75
val hyps_file = File.tmp_path (Path.basic "hyps");
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    76
val prob_file = File.tmp_path (Path.basic "prob");
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    77
val dummy_tac = PRIMITIVE(fn thm => thm );
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    78
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    79
 
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    80
(**** for Isabelle/ML interface  ****)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    81
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
    82
fun is_proof_char ch = ((33 <= (ord ch)) andalso ((ord ch ) <= 126) andalso (not ((ord ch ) = 63))) orelse (ch = " ")
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    83
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
    84
fun proofstring x = let val exp = explode x 
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
    85
                    in
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
    86
                        List.filter (is_proof_char ) exp
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
    87
                    end
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
    88
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    89
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
    90
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    91
(*
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    92
fun call_atp_tac thms n = (tptp_inputs thms ; dummy_tac);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    93
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    94
*)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    95
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    96
(**** For running in Isar ****)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    97
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    98
(* same function as that in res_axioms.ML *)
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    99
fun repeat_RS thm1 thm2 =
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   100
    let val thm1' =  thm1 RS thm2 handle THM _ => thm1
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   101
    in
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   102
	if eq_thm(thm1,thm1') then thm1' else (repeat_RS thm1' thm2)
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   103
    end;
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   104
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   105
(* a special version of repeat_RS *)
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   106
fun repeat_someI_ex thm = repeat_RS thm someI_ex;
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   107
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   108
(*********************************************************************)
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   109
(* convert clauses from "assume" to conjecture. write to file "hyps" *)
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   110
(* hypotheses of the goal currently being proved                     *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   111
(*********************************************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   112
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   113
fun isar_atp_h thms =
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   114
        
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   115
    let val prems = map (skolemize o make_nnf o ObjectLogic.atomize_thm) thms
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   116
        val prems' = map repeat_someI_ex prems
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   117
        val prems'' = make_clauses prems'
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   118
        val prems''' = ResAxioms.rm_Eps [] prems''
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   119
        val clss = map ResClause.make_conjecture_clause prems'''
15774
9df37a0e935d more tidying of libraries in Reconstruction
paulson
parents: 15736
diff changeset
   120
	val (tptp_clss,tfree_litss) = ListPair.unzip (map ResClause.clause2tptp clss) 
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   121
	val tfree_lits = ResLib.flat_noDup tfree_litss
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   122
	val tfree_clss = map ResClause.tfree_clause tfree_lits 
16259
aed1a8ae4b23 File.platform_path;
wenzelm
parents: 16185
diff changeset
   123
        val hypsfile = File.platform_path hyps_file
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   124
	val out = TextIO.openOut(hypsfile)
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   125
    in
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   126
	((ResLib.writeln_strs out (tfree_clss @ tptp_clss); TextIO.closeOut out; if !trace_res then (warning hypsfile) else ());tfree_lits) 
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   127
    end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   128
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   129
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   130
(*********************************************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   131
(* write out a subgoal as tptp clauses to the file "probN"           *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   132
(* where N is the number of this subgoal                             *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   133
(*********************************************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   134
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   135
fun tptp_inputs_tfrees thms n tfrees = 
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   136
    let val _ = (warning ("in tptp_inputs_tfrees 0"))
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   137
        val clss = map (ResClause.make_conjecture_clause_thm) thms
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   138
         val _ = (warning ("in tptp_inputs_tfrees 1"))
15774
9df37a0e935d more tidying of libraries in Reconstruction
paulson
parents: 15736
diff changeset
   139
	val (tptp_clss,tfree_litss) = ListPair.unzip (map ResClause.clause2tptp clss)
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   140
        val _ = (warning ("in tptp_inputs_tfrees 2"))
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   141
	val tfree_clss = map ResClause.tfree_clause ((ResLib.flat_noDup tfree_litss) \\ tfrees) 
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   142
         val _ = (warning ("in tptp_inputs_tfrees 3"))
16259
aed1a8ae4b23 File.platform_path;
wenzelm
parents: 16185
diff changeset
   143
        val probfile = (File.platform_path prob_file) ^ "_" ^ (string_of_int n)
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   144
	val out = TextIO.openOut(probfile)
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   145
    in
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   146
	(ResLib.writeln_strs out (tfree_clss @ tptp_clss); TextIO.closeOut out; (if !trace_res then (warning probfile) else ()))
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   147
    end;
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
   148
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   149
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   150
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   151
(*********************************************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   152
(* call SPASS with settings and problem file for the current subgoal *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   153
(* should be modified to allow other provers to be called            *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   154
(*********************************************************************)
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   155
(* now passing in list of skolemized thms and list of sgterms to go with them *)
16515
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   156
fun call_resolve_tac  (thms: Thm.thm list list)  sign (sg_terms:  Term.term list) (childin, childout,pid) n  = 
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   157
 let
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   158
   val axfile = (File.platform_path axiom_file)
16475
8f3ba52a7937 using TPTP2X_HOME; indentation, etc
paulson
parents: 16357
diff changeset
   159
   val hypsfile = (File.platform_path hyps_file)
8f3ba52a7937 using TPTP2X_HOME; indentation, etc
paulson
parents: 16357
diff changeset
   160
   val clasimpfile = (File.platform_path clasimp_file)
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   161
     
16475
8f3ba52a7937 using TPTP2X_HOME; indentation, etc
paulson
parents: 16357
diff changeset
   162
   fun make_atp_list [] sign n = []
8f3ba52a7937 using TPTP2X_HOME; indentation, etc
paulson
parents: 16357
diff changeset
   163
   |   make_atp_list ((sko_thm, sg_term)::xs) sign  n = 
16515
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   164
     let 
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   165
	val skothmstr = Meson.concat_with_and (map string_of_thm sko_thm) 
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   166
	val thmproofstr = proofstring ( skothmstr)
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   167
	val no_returns =List.filter   not_newline ( thmproofstr)
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   168
	val thmstr = implode no_returns
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   169
	val _ = warning ("thmstring in make_atp_lists is: "^thmstr)
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   170
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   171
	val sgstr = Sign.string_of_term sign sg_term 
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   172
	val goalproofstring = proofstring sgstr
16089
9169bdf930f8 trying to set up portable calling sequences for SPASS and tptp2X
paulson
parents: 16061
diff changeset
   173
	val no_returns =List.filter not_newline ( goalproofstring)
9169bdf930f8 trying to set up portable calling sequences for SPASS and tptp2X
paulson
parents: 16061
diff changeset
   174
	val goalstring = implode no_returns
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   175
	val _ = warning ("goalstring in make_atp_lists is: "^goalstring)
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   176
16259
aed1a8ae4b23 File.platform_path;
wenzelm
parents: 16185
diff changeset
   177
	val probfile = (File.platform_path prob_file) ^ "_" ^ (string_of_int n)
16515
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   178
	val _ = (warning ("prob file in cal_res_tac is: "^probfile))      
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   179
     in
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   180
       if !SpassComm.spass 
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   181
       then 
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   182
         let val _ = ResLib.helper_path "SPASS_HOME" "SPASS"
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   183
         in  (*We've checked that SPASS is there for ATP/spassshell to run.*)
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   184
           if !full_spass 
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   185
            then (*Allow SPASS to run in Auto mode, using all its inference rules*)
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   186
   		([("spass", thmstr, goalstring (*,spass_home*), 
16515
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   187
	    	 getenv "ISABELLE_HOME" ^ "/src/HOL/Tools/ATP/spassshell",  
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   188
	     	"-DocProof%-TimeLimit=60", 
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   189
	     	clasimpfile, axfile, hypsfile, probfile)]@(make_atp_list xs sign (n+1)))
16515
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   190
  	    else (*Restrict SPASS to a small set of rules that we can parse*)
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   191
   		([("spass", thmstr, goalstring (*,spass_home*), 
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   192
	     	getenv "ISABELLE_HOME" ^"/src/HOL/Tools/ATP/spassshell"(* (getenv "SPASS_HOME")^"/SPASS"*),  
16478
d0a1f6231e2f Added VampCommunication.ML.
quigley
parents: 16475
diff changeset
   193
	     	"-Auto=0%-IORe%-IOFc%-RTaut%-RFSub%-RBSub%-DocProof%-TimeLimit=60", 
d0a1f6231e2f Added VampCommunication.ML.
quigley
parents: 16475
diff changeset
   194
	     	clasimpfile, axfile, hypsfile, probfile)]@(make_atp_list xs sign (n+1)))
16515
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   195
	 end
16478
d0a1f6231e2f Added VampCommunication.ML.
quigley
parents: 16475
diff changeset
   196
     else
16515
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   197
       let val vampire = ResLib.helper_path "VAMPIRE_HOME" "vkernel"
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   198
       in  
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   199
	   ([("vampire", thmstr, goalstring, vampire, "-t 300%-m 100000", 
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   200
	      clasimpfile, axfile, hypsfile, probfile)] @ 
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   201
	    (make_atp_list xs sign (n+1)))
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   202
       end
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   203
     end
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
   204
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   205
val thms_goals = ListPair.zip( thms, sg_terms) 
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   206
val atp_list = make_atp_list  thms_goals sign 1
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   207
in
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   208
	Watcher.callResProvers(childout,atp_list);
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   209
        warning("Sent commands to watcher!");
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   210
  	dummy_tac
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   211
 end
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   212
(*
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   213
  val axfile = (File.platform_path axiom_file)
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   214
    val hypsfile = (File.platform_path hyps_file)
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   215
     val clasimpfile = (File.platform_path  clasimp_file)
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   216
    val spass_home = getenv "SPASS_HOME"
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   217
 val probfile = (File.platform_path prob_file) ^ "_" ^ (string_of_int 1)
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   218
    
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   219
 val (clause_arr, num_of_clauses) = ResClasimp.write_out_clasimp_small (File.platform_path clasimp_file) ;
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   220
 val (childin,childout,pid) = Watcher.createWatcher(mp, clause_arr, num_of_clauses);
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   221
Watcher.callResProvers(childout, [("spass", "thmstring", "goalstr", spass_home,  "-DocProof", clasimpfile, axfile, hypsfile,probfile)]);
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   222
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   223
Watcher.callResProvers(childout, [("spass", "thmstring", "goalstr", "/homes/clq20/spassshell",  "", clasimpfile, axfile, hypsfile,probfile)]);
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   224
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   225
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   226
*)
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   227
(**********************************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   228
(* write out the current subgoal as a tptp file, probN,   *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   229
(* then call dummy_tac - should be call_res_tac           *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   230
(**********************************************************)
15679
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   231
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   232
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   233
fun get_sko_thms tfrees sign sg_terms (childin, childout,pid)  thm n sko_thms  =
16515
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   234
    if n=0 then 
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   235
	     (call_resolve_tac  (rev sko_thms) sign  sg_terms (childin, childout, pid) (List.length sg_terms);dummy_tac thm)
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   236
     else
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   237
	
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   238
     ( SELECT_GOAL
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   239
	     (EVERY1 [rtac ccontr,atomize_tac, skolemize_tac, 
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   240
	      METAHYPS(fn negs => (tptp_inputs_tfrees (make_clauses negs) n tfrees;
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   241
				 get_sko_thms  tfrees sign sg_terms (childin, childout, pid) thm  (n -1) (negs::sko_thms);dummy_tac))]) n thm )
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   242
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   243
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   244
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   245
(**********************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   246
(* recursively call atp_tac_g on all subgoals *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   247
(* sg_term is the nth subgoal as a term - used*)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   248
(* in proof reconstruction                    *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   249
(**********************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   250
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   251
fun isar_atp_goal' thm n tfree_lits  (childin, childout, pid) = 
16515
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   252
    let val  prems = prems_of thm 
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   253
      (*val sg_term = get_nth k prems*)
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   254
	val sign = sign_of_thm thm
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   255
	val thmstring = string_of_thm thm
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   256
    in   
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   257
	warning("in isar_atp_goal'");
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   258
	warning("thmstring in isar_atp_goal': "^thmstring);
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   259
	(* go and call callResProvers with this subgoal *)
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   260
	(* isar_atp_g tfree_lits  sg_term (childin, childout, pid) k thm; *)
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   261
	(* recursive call to pick up the remaining subgoals *)
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   262
	(* isar_atp_goal' thm (k+1) n tfree_lits  (childin, childout, pid) *)
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   263
	get_sko_thms tfree_lits sign  prems (childin, childout, pid) thm n []
7896ea4f3a87 VAMPIRE_HOME, helper_path and various stylistic tweaks
paulson
parents: 16478
diff changeset
   264
    end ;
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   265
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   266
(*fun isar_atp_goal thm n_subgoals tfree_lits   (childin, childout, pid) = 
15779
aed221aff642 Removed remaining references to Main.thy in reconstruction code.
quigley
parents: 15774
diff changeset
   267
              (if (!debug) then warning (string_of_thm thm) 
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   268
               else (isar_atp_goal' thm n_subgoals tfree_lits  (childin, childout, pid) ));*)
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   269
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   270
(**************************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   271
(* convert clauses from "assume" to conjecture.   *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   272
(* i.e. apply make_clauses and then get tptp for  *)
15679
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   273
(* any hypotheses in the goal produced by assume  *)
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   274
(* statements;                                    *)
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   275
(* write to file "hyps"                           *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   276
(**************************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   277
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   278
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   279
fun isar_atp_aux thms thm n_subgoals  (childin, childout, pid) = 
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   280
    let val tfree_lits = isar_atp_h thms 
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   281
    in
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   282
	(warning("in isar_atp_aux"));
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   283
         isar_atp_goal' thm n_subgoals tfree_lits   (childin, childout, pid)
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   284
    end;
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   285
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   286
(******************************************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   287
(* called in Isar automatically                                   *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   288
(* writes out the current clasimpset to a tptp file               *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   289
(* passes all subgoals on to isar_atp_aux for further processing  *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   290
(* turns off xsymbol at start of function, restoring it at end    *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   291
(******************************************************************)
15779
aed221aff642 Removed remaining references to Main.thy in reconstruction code.
quigley
parents: 15774
diff changeset
   292
(*FIX changed to clasimp_file *)
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   293
fun isar_atp' (ctxt,thms, thm) =
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   294
    let 
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   295
	val _=  (print_mode := (Library.gen_rems (op =) (! print_mode, ["xsymbols", "symbols"])))
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   296
        val _= (warning ("in isar_atp'"))
15779
aed221aff642 Removed remaining references to Main.thy in reconstruction code.
quigley
parents: 15774
diff changeset
   297
        val prems  = prems_of thm
15697
681bcb7f0389 removal of Main and other tidying up
paulson
parents: 15682
diff changeset
   298
        val sign = sign_of_thm thm
15779
aed221aff642 Removed remaining references to Main.thy in reconstruction code.
quigley
parents: 15774
diff changeset
   299
        val thms_string = Meson.concat_with_and (map string_of_thm thms) 
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   300
        val thmstring = string_of_thm thm
15779
aed221aff642 Removed remaining references to Main.thy in reconstruction code.
quigley
parents: 15774
diff changeset
   301
        val prems_string = Meson.concat_with_and (map (Sign.string_of_term sign) prems) 
15919
b30a35432f5a Replaced reference to SPASS with general one - set SPASS_HOME in settings file.
quigley
parents: 15782
diff changeset
   302
        
b30a35432f5a Replaced reference to SPASS with general one - set SPASS_HOME in settings file.
quigley
parents: 15782
diff changeset
   303
	(* set up variables for writing out the clasimps to a tptp file *)
16172
629ba53a072f small tweaks; also now write_out_clasimp takes the current theory as argument
paulson
parents: 16156
diff changeset
   304
	val (clause_arr, num_of_clauses) =
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   305
	    ResClasimp.write_out_clasimp(*_small*) (File.platform_path clasimp_file) 
16172
629ba53a072f small tweaks; also now write_out_clasimp takes the current theory as argument
paulson
parents: 16156
diff changeset
   306
	                                 (ProofContext.theory_of ctxt)
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   307
        val _ = (warning ("clasimp_file is this: "^(File.platform_path clasimp_file)) )  
15779
aed221aff642 Removed remaining references to Main.thy in reconstruction code.
quigley
parents: 15774
diff changeset
   308
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   309
        (* cq: add write out clasimps to file *)
15919
b30a35432f5a Replaced reference to SPASS with general one - set SPASS_HOME in settings file.
quigley
parents: 15782
diff changeset
   310
        (* cq:create watcher and pass to isar_atp_aux *)
b30a35432f5a Replaced reference to SPASS with general one - set SPASS_HOME in settings file.
quigley
parents: 15782
diff changeset
   311
        (* tracing *) 
b30a35432f5a Replaced reference to SPASS with general one - set SPASS_HOME in settings file.
quigley
parents: 15782
diff changeset
   312
        (*val tenth_ax = fst( Array.sub(clause_arr, 1))  
b30a35432f5a Replaced reference to SPASS with general one - set SPASS_HOME in settings file.
quigley
parents: 15782
diff changeset
   313
         val tenth_ax_thms = Recon_Transfer.memo_find_clause (tenth_ax, clause_tab)
b30a35432f5a Replaced reference to SPASS with general one - set SPASS_HOME in settings file.
quigley
parents: 15782
diff changeset
   314
         val clause_str = Meson.concat_with_and (map string_of_thm tenth_ax_thms)
b30a35432f5a Replaced reference to SPASS with general one - set SPASS_HOME in settings file.
quigley
parents: 15782
diff changeset
   315
         val _ = (warning ("tenth axiom in array is: "^tenth_ax))         
b30a35432f5a Replaced reference to SPASS with general one - set SPASS_HOME in settings file.
quigley
parents: 15782
diff changeset
   316
         val _ = (warning ("tenth axiom in table is: "^clause_str))         
b30a35432f5a Replaced reference to SPASS with general one - set SPASS_HOME in settings file.
quigley
parents: 15782
diff changeset
   317
                 
b30a35432f5a Replaced reference to SPASS with general one - set SPASS_HOME in settings file.
quigley
parents: 15782
diff changeset
   318
         val _ = (warning ("num_of_clauses is: "^(string_of_int (num_of_clauses))) )     
b30a35432f5a Replaced reference to SPASS with general one - set SPASS_HOME in settings file.
quigley
parents: 15782
diff changeset
   319
         *)             
b30a35432f5a Replaced reference to SPASS with general one - set SPASS_HOME in settings file.
quigley
parents: 15782
diff changeset
   320
        
b30a35432f5a Replaced reference to SPASS with general one - set SPASS_HOME in settings file.
quigley
parents: 15782
diff changeset
   321
        val (childin,childout,pid) = Watcher.createWatcher(thm,clause_arr, num_of_clauses)
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   322
        val pidstring = string_of_int(Word.toInt (Word.fromLargeWord ( Posix.Process.pidToWord pid )))
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   323
    in
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   324
	case prems of [] => () 
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   325
		    | _ => ((warning ("initial thms: "^thms_string)); 
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   326
                           (warning ("initial thm: "^thmstring));
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   327
                           (warning ("subgoals: "^prems_string));
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   328
                           (warning ("pid: "^ pidstring))); 
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   329
                            isar_atp_aux thms thm (length prems) (childin, childout, pid) ;
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   330
                           (warning("turning xsymbol back on!"));
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   331
                           print_mode := (["xsymbols", "symbols"] @ ! print_mode)
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   332
    end;
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   333
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
   334
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   335
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   336
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   337
local
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   338
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   339
fun get_thms_cs claset =
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   340
    let val clsset = rep_cs claset
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   341
	val safeEs = #safeEs clsset
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   342
	val safeIs = #safeIs clsset
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   343
	val hazEs = #hazEs clsset
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   344
	val hazIs = #hazIs clsset
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   345
    in
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   346
	safeEs @ safeIs @ hazEs @ hazIs
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   347
    end;
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   348
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   349
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   350
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   351
fun append_name name [] _ = []
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   352
  | append_name name (thm::thms) k = (Thm.name_thm ((name ^ "_" ^ (string_of_int k)),thm)) :: (append_name name thms (k+1));
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   353
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   354
fun append_names (name::names) (thms::thmss) =
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   355
    let val thms' = append_name name thms 0
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   356
    in
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   357
	thms'::(append_names names thmss)
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   358
    end;
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   359
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   360
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   361
fun get_thms_ss [] = []
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   362
  | get_thms_ss thms =
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   363
    let val names = map Thm.name_of_thm thms 
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   364
        val thms' = map (mksimps mksimps_pairs) thms
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   365
        val thms'' = append_names names thms'
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   366
    in
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   367
	ResLib.flat_noDup thms''
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   368
    end;
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   369
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   370
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   371
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   372
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   373
in
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   374
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
   375
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   376
(* convert locally declared rules to axiom clauses *)
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   377
(* write axiom clauses to ax_file *)
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   378
(* what about clasimpset - it should already be in the ax file - perhaps append to ax file rather than just *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   379
(* write out ? Or keep as a separate file and then cat them all together in the watcher, like we do with the *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   380
(*claset file and prob file*)
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   381
(* FIX*)
16039
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 15919
diff changeset
   382
(*fun isar_local_thms (delta_cs, delta_ss_thms) =
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   383
    let val thms_cs = get_thms_cs delta_cs
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   384
	val thms_ss = get_thms_ss delta_ss_thms
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   385
	val thms_clauses = ResLib.flat_noDup (map ResAxioms.clausify_axiom (thms_cs @ thms_ss))
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   386
	val clauses_strs = ResLib.flat_noDup (map ResClause.tptp_clause thms_clauses) (*string list*)
16259
aed1a8ae4b23 File.platform_path;
wenzelm
parents: 16185
diff changeset
   387
	val ax_file = File.platform_path axiom_file
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   388
	val out = TextIO.openOut ax_file
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   389
    in
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   390
	(ResLib.writeln_strs out clauses_strs; (warning ("axiom file is: "^ax_file));TextIO.closeOut out)
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   391
    end;
16039
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 15919
diff changeset
   392
*)
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   393
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   394
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   395
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   396
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   397
(* called in Isar automatically *)
15679
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   398
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   399
fun isar_atp (ctxt,thm) =
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   400
    let val prems = ProofContext.prems_of ctxt
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   401
        val d_cs = Classical.get_delta_claset ctxt 
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   402
        val d_ss_thms = Simplifier.get_delta_simpset ctxt
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   403
        val thmstring = string_of_thm thm
15679
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   404
        val sg_prems = prems_of thm
15697
681bcb7f0389 removal of Main and other tidying up
paulson
parents: 15682
diff changeset
   405
        val sign = sign_of_thm thm
15679
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   406
        val prem_no = length sg_prems
15779
aed221aff642 Removed remaining references to Main.thy in reconstruction code.
quigley
parents: 15774
diff changeset
   407
        val prems_string = Meson.concat_with_and (map (Sign.string_of_term sign) sg_prems) 
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   408
    in
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   409
          (warning ("initial thm in isar_atp: "^thmstring));
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   410
          (warning ("subgoals in isar_atp: "^prems_string));
15779
aed221aff642 Removed remaining references to Main.thy in reconstruction code.
quigley
parents: 15774
diff changeset
   411
    	  (warning ("number of subgoals in isar_atp: "^(string_of_int prem_no)));
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   412
          (*isar_local_thms (d_cs,d_ss_thms);*)
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   413
           isar_atp' (ctxt,prems, thm)
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   414
    end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   415
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   416
end
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   417
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   418
16357
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   419
f1275d2a1dee All subgoals sent to the watcher at once now.
quigley
parents: 16259
diff changeset
   420
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   421
end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   422
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   423
Proof.atp_hook := ResAtp.isar_atp;