src/HOL/Tools/res_atp.ML
author paulson
Tue, 12 Apr 2005 11:08:25 +0200
changeset 15700 970e0293dfb3
parent 15697 681bcb7f0389
child 15736 1bb0399a9517
permissions -rw-r--r--
tweaks mainly to achieve sml/nj compatibility
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
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    11
signature RES_ATP = 
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    12
sig
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    13
val trace_res : bool ref
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    14
val axiom_file : Path.T
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    15
val hyps_file : Path.T
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    16
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
    17
val prob_file : Path.T;
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    18
(*val atp_ax_tac : Thm.thm list -> int -> Tactical.tactic*)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    19
(*val atp_tac : int -> Tactical.tactic*)
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    20
val debug: bool ref
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    21
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    22
end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    23
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    24
structure ResAtp : RES_ATP =
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    25
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    26
struct
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    27
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    28
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    29
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    30
(* used for debug *)
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    31
val debug = ref false;
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
    32
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    33
fun debug_tac tac = (warning "testing";tac);
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    34
(* default value is false *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    35
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    36
val trace_res = ref false;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    37
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    38
val skolem_tac = skolemize_tac;
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    39
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    40
val num_of_clauses = ref 0;
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    41
val clause_arr = Array.array(3500, ("empty", 0));
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    42
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    43
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    44
val atomize_tac =
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    45
    SUBGOAL
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    46
     (fn (prop,_) =>
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    47
	 let val ts = Logic.strip_assums_hyp prop
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    48
	 in EVERY1 
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    49
		[METAHYPS
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    50
		     (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
    51
	  REPEAT_DETERM_N (length ts) o (etac thin_rl)]
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    52
     end);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    53
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    54
(* temporarily use these files, which will be loaded by Vampire *)
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    55
val file_id_num =ref 0;
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    56
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    57
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
    58
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    59
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    60
val axiom_file = File.tmp_path (Path.basic "axioms");
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    61
val clasimp_file = File.tmp_path (Path.basic "clasimp");
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    62
val hyps_file = File.tmp_path (Path.basic "hyps");
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    63
val prob_file = File.tmp_path (Path.basic "prob");
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    64
val dummy_tac = PRIMITIVE(fn thm => thm );
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    65
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    66
 
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    67
fun concat_with_and [] str = str
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    68
|   concat_with_and (x::[]) str = str^" ("^x^")"
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    69
|   concat_with_and (x::xs) str = (concat_with_and xs (str^"("^x^")"^" & "))
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    70
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    71
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    72
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    73
(**** for Isabelle/ML interface  ****)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    74
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    75
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
    76
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    77
fun proofstring x = let val exp = explode x 
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    78
                    in
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    79
                        List.filter (is_proof_char ) exp
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    80
                    end
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    81
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    82
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
    83
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    84
(*
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    85
fun call_atp_tac thms n = (tptp_inputs thms ; dummy_tac);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    86
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
    87
*)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    88
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    89
(**** For running in Isar ****)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    90
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    91
(* same function as that in res_axioms.ML *)
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    92
fun repeat_RS thm1 thm2 =
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    93
    let val thm1' =  thm1 RS thm2 handle THM _ => thm1
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    94
    in
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    95
	if eq_thm(thm1,thm1') then thm1' else (repeat_RS thm1' thm2)
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    96
    end;
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    97
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    98
(* a special version of repeat_RS *)
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
    99
fun repeat_someI_ex thm = repeat_RS thm someI_ex;
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   100
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   101
(*********************************************************************)
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   102
(* convert clauses from "assume" to conjecture. write to file "hyps" *)
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   103
(* hypotheses of the goal currently being proved                     *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   104
(*********************************************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   105
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   106
fun isar_atp_h thms =
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   107
        
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   108
    let val prems = map (skolemize o make_nnf o ObjectLogic.atomize_thm) thms
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   109
        val prems' = map repeat_someI_ex prems
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   110
        val prems'' = make_clauses prems'
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   111
        val prems''' = ResAxioms.rm_Eps [] prems''
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   112
        val clss = map ResClause.make_conjecture_clause prems'''
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   113
	val (tptp_clss,tfree_litss) = ResLib.unzip (map ResClause.clause2tptp clss) 
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   114
	val tfree_lits = ResLib.flat_noDup tfree_litss
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   115
	val tfree_clss = map ResClause.tfree_clause tfree_lits 
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   116
        val hypsfile = File.sysify_path hyps_file
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   117
	val out = TextIO.openOut(hypsfile)
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   118
    in
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   119
	((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
   120
    end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   121
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   122
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   123
(*********************************************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   124
(* write out a subgoal as tptp clauses to the file "probN"           *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   125
(* where N is the number of this subgoal                             *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   126
(*********************************************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   127
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   128
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
   129
    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
   130
        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
   131
         val _ = (warning ("in tptp_inputs_tfrees 1"))
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   132
	val (tptp_clss,tfree_litss) = ResLib.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
   133
        val _ = (warning ("in tptp_inputs_tfrees 2"))
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   134
	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
   135
         val _ = (warning ("in tptp_inputs_tfrees 3"))
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   136
        val probfile = (File.sysify_path prob_file) ^ "_" ^ (string_of_int n)
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   137
	val out = TextIO.openOut(probfile)
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   138
    in
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   139
	(ResLib.writeln_strs out (tfree_clss @ tptp_clss); TextIO.closeOut out; (if !trace_res then (warning probfile) else ()))
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   140
    end;
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
   141
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   142
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   143
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   144
(*********************************************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   145
(* call SPASS with settings and problem file for the current subgoal *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   146
(* should be modified to allow other provers to be called            *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   147
(*********************************************************************)
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   148
15697
681bcb7f0389 removal of Main and other tidying up
paulson
parents: 15682
diff changeset
   149
fun call_resolve_tac sign thms sg_term (childin, childout,pid) n  = let
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   150
                             val thmstring = concat_with_and (map string_of_thm thms) ""
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   151
                             val thm_no = length thms
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   152
                             val _ = warning ("number of thms is : "^(string_of_int thm_no))
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   153
                             val _ = warning ("thmstring in call_res is: "^thmstring)
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   154
15697
681bcb7f0389 removal of Main and other tidying up
paulson
parents: 15682
diff changeset
   155
                             val goalstr = Sign.string_of_term sign sg_term 
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   156
                             val goalproofstring = proofstring goalstr
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   157
                             val no_returns =List.filter not_newline ( goalproofstring)
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   158
                             val goalstring = implode no_returns
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   159
                             val _ = warning ("goalstring in call_res is: "^goalstring)
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   160
        
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   161
                             (*val prob_file =File.tmp_path (Path.basic newprobfile); *)
15653
3549ff7158f3 Updated to add watcher code.
quigley
parents: 15644
diff changeset
   162
                             (*val _ =( warning ("calling make_clauses "))
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   163
                             val clauses = make_clauses thms
15653
3549ff7158f3 Updated to add watcher code.
quigley
parents: 15644
diff changeset
   164
                             val _ =( warning ("called make_clauses "))*)
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   165
                             (*val _ = tptp_inputs clauses prob_file*)
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   166
                             val thmstring = concat_with_and (map string_of_thm thms) ""
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   167
                           
15697
681bcb7f0389 removal of Main and other tidying up
paulson
parents: 15682
diff changeset
   168
                             val goalstr = Sign.string_of_term sign sg_term 
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   169
                             val goalproofstring = proofstring goalstr
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   170
                             val no_returns =List.filter not_newline ( goalproofstring)
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   171
                             val goalstring = implode no_returns
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   172
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   173
                             val thmproofstring = proofstring ( thmstring)
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   174
                             val no_returns =List.filter   not_newline ( thmproofstring)
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   175
                             val thmstr = implode no_returns
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   176
                            
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   177
                             val probfile = (File.sysify_path prob_file) ^ "_" ^ (string_of_int n)
15679
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   178
                             val outfile = TextIO.openOut(File.sysify_path(File.tmp_path (Path.basic "hellofile")))
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   179
                             val _ = TextIO.output(outfile, "prob file path is "^probfile^" thmstring is "^thmstr^" goalstring is "^goalstring);
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   180
                             val _ = TextIO.flushOut outfile;
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   181
                             val _ =  TextIO.closeOut outfile
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   182
                          in
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   183
                           (* without paramodulation *)
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   184
                           (warning ("goalstring in call_res_tac is: "^goalstring));
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   185
                           (warning ("prob file in cal_res_tac is: "^probfile));
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   186
                            Watcher.callResProvers(childout,
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   187
                            [("spass",thmstr,goalstring,"/homes/clq20/bin/SPASS",  
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   188
                             "-FullRed=0%-Auto=0%-ISRe%-ISFc%-RTaut%-RFSub%-RBSub%-DocProof", 
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   189
                             probfile)]);
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   190
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   191
                           (* with paramodulation *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   192
                           (*Watcher.callResProvers(childout,
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   193
                                  [("spass",thmstr,goalstring,"/homes/clq20/bin/SPASS",
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   194
                                  "-FullRed=0%-ISPm=1%-Split=0%-PObv=0%-DocProof", 
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   195
                                    prob_path)]); *)
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   196
                          (* Watcher.callResProvers(childout,
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   197
                           [("spass",thmstr,goalstring,"/homes/clq20/bin/SPASS", 
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   198
                           "-DocProof",  prob_path)]);*)
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   199
                           dummy_tac
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   200
                         end
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
   201
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   202
(************************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   203
(* pass in subgoal as a term and watcher info   *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   204
(* process subgoal into skolemized, negated form*)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   205
(* then call call_resolve_tac to send to ATP    *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   206
(************************************************)
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   207
(*
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   208
fun resolve_tac sg_term  (childin, childout,pid) = 
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   209
   let val _ = warning ("in resolve_tac ")
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   210
   in
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   211
   SELECT_GOAL
15653
3549ff7158f3 Updated to add watcher code.
quigley
parents: 15644
diff changeset
   212
  (EVERY1 [rtac ccontr,atomize_tac,skolemize_tac,  METAHYPS(fn negs => (warning ("calling call_resolve_tac next ");dummy_tac (*call_resolve_tac negs sg_term (childin, childout,pid)*)))])
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   213
   end;
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
   214
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   215
*)
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   216
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
   217
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   218
(* Need to replace call_atp_tac_tfrees with call res_provers as it's the dummy one *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   219
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   220
(**********************************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   221
(* write out the current subgoal as a tptp file, probN,   *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   222
(* then call dummy_tac - should be call_res_tac           *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   223
(**********************************************************)
15653
3549ff7158f3 Updated to add watcher code.
quigley
parents: 15644
diff changeset
   224
(* should call call_res_tac here, not resolve_tac *)
3549ff7158f3 Updated to add watcher code.
quigley
parents: 15644
diff changeset
   225
(* if we take tptpinputs out it gets into call_res_tac then falls over as usual after printing out goalstring. but if we leave it in it falls over here *)
15679
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   226
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   227
(* dummy tac vs.  Doesn't call resolve_tac *)
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   228
15697
681bcb7f0389 removal of Main and other tidying up
paulson
parents: 15682
diff changeset
   229
fun call_atp_tac_tfrees sign thms n tfrees sg_term (childin, childout,pid) = 
681bcb7f0389 removal of Main and other tidying up
paulson
parents: 15682
diff changeset
   230
                                         ( warning("ths for tptp: " ^ (concat_with_and (map string_of_thm thms) ""));
681bcb7f0389 removal of Main and other tidying up
paulson
parents: 15682
diff changeset
   231
                                           warning("in call_atp_tac_tfrees");
15679
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   232
                                           
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   233
                                           tptp_inputs_tfrees (make_clauses thms) n tfrees;
15697
681bcb7f0389 removal of Main and other tidying up
paulson
parents: 15682
diff changeset
   234
                                           call_resolve_tac sign thms sg_term (childin, childout, pid) n;
15679
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   235
  					   dummy_tac);
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   236
15697
681bcb7f0389 removal of Main and other tidying up
paulson
parents: 15682
diff changeset
   237
fun atp_tac_tfrees tfrees sg_term (childin, childout,pid)  n st = 
681bcb7f0389 removal of Main and other tidying up
paulson
parents: 15682
diff changeset
   238
let val sign = sign_of_thm st
681bcb7f0389 removal of Main and other tidying up
paulson
parents: 15682
diff changeset
   239
    val _ = warning ("in atp_tac_tfrees ")
681bcb7f0389 removal of Main and other tidying up
paulson
parents: 15682
diff changeset
   240
    val _ = warning ("sg_term :" ^ (Sign.string_of_term sign sg_term))
15679
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   241
   
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   242
   in
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   243
SELECT_GOAL
15682
09a7b8909c4d removed bad code
paulson
parents: 15679
diff changeset
   244
  (EVERY1 [rtac ccontr,atomize_tac, skolemize_tac, 
15697
681bcb7f0389 removal of Main and other tidying up
paulson
parents: 15682
diff changeset
   245
  METAHYPS(fn negs => (warning("calling  call_atp_tac_tfrees with negs"
681bcb7f0389 removal of Main and other tidying up
paulson
parents: 15682
diff changeset
   246
                               ^ (concat_with_and (map string_of_thm negs) ""));
681bcb7f0389 removal of Main and other tidying up
paulson
parents: 15682
diff changeset
   247
           call_atp_tac_tfrees sign negs n tfrees sg_term (childin, childout,pid) ))]) n st
15679
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   248
end;
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   249
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   250
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   251
fun isar_atp_g tfrees sg_term (childin, childout, pid) n = 
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   252
                                        
15679
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   253
((warning("in isar_atp_g"));atp_tac_tfrees tfrees sg_term (childin, childout, pid) n);
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   254
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   255
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   256
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   257
(**********************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   258
(* recursively call atp_tac_g on all subgoals *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   259
(* sg_term is the nth subgoal as a term - used*)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   260
(* in proof reconstruction                    *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   261
(**********************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   262
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   263
fun isar_atp_goal' thm k n tfree_lits  (childin, childout, pid) = 
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   264
                  	if (k > n) 
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   265
                        then () 
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   266
	  		else 
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   267
                           (let val  prems = prems_of thm 
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   268
                                val sg_term = get_nth n prems
15679
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   269
                                val thmstring = string_of_thm thm
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   270
                            in   
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   271
                                 
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   272
                		(warning("in isar_atp_goal'"));
15679
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   273
                                (warning("thmstring in isar_atp_goal: "^thmstring));
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   274
 				 isar_atp_g tfree_lits  sg_term (childin, childout, pid) k thm; 
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   275
                                 isar_atp_goal' thm (k+1) n tfree_lits  (childin, childout, pid) 
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   276
                            end);
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_goal thm n_subgoals tfree_lits   (childin, childout, pid) = (if (!debug) then warning (string_of_thm thm) else (isar_atp_goal' thm 1 n_subgoals tfree_lits  (childin, childout, pid) ));
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   280
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   281
(**************************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   282
(* convert clauses from "assume" to conjecture.   *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   283
(* 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
   284
(* any hypotheses in the goal produced by assume  *)
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   285
(* statements;                                    *)
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   286
(* write to file "hyps"                           *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   287
(**************************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   288
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   289
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   290
fun isar_atp_aux thms thm n_subgoals  (childin, childout, pid) = 
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   291
    let val tfree_lits = isar_atp_h thms 
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   292
    in
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   293
	(warning("in isar_atp_aux"));isar_atp_goal thm n_subgoals tfree_lits   (childin, childout, pid)
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   294
    end;
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   295
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   296
(******************************************************************)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   297
(* called in Isar automatically                                   *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   298
(* writes out the current clasimpset to a tptp file               *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   299
(* passes all subgoals on to isar_atp_aux for further processing  *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   300
(* turns off xsymbol at start of function, restoring it at end    *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   301
(******************************************************************)
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
   302
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   303
fun isar_atp' (thms, thm) =
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   304
    let val _=  (print_mode := (Library.gen_rems (op =) (! print_mode, ["xsymbols", "symbols"])))
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   305
        val _= (warning ("in isar_atp'"))
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   306
        val prems = prems_of thm
15697
681bcb7f0389 removal of Main and other tidying up
paulson
parents: 15682
diff changeset
   307
        val sign = sign_of_thm thm
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   308
        val thms_string =concat_with_and (map  string_of_thm thms) ""
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   309
        val thmstring = string_of_thm thm
15697
681bcb7f0389 removal of Main and other tidying up
paulson
parents: 15682
diff changeset
   310
        val prems_string = concat_with_and (map (Sign.string_of_term sign) prems) ""
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   311
        (* set up variables for writing out the clasimps to a tptp file *)
15700
970e0293dfb3 tweaks mainly to achieve sml/nj compatibility
paulson
parents: 15697
diff changeset
   312
        val _ = write_out_clasimp (File.sysify_path axiom_file)
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   313
        (* cq: add write out clasimps to file *)
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   314
        (* cq:create watcher and pass to isar_atp_aux *)                    
15657
bd946fbc7c2b Current version of res_atp.ML - causes an error when I run it. C.Q.
quigley
parents: 15653
diff changeset
   315
        val (childin,childout,pid) = Watcher.createWatcher()
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   316
        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
   317
    in
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   318
	case prems of [] => () 
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   319
		    | _ => ((warning ("initial thms: "^thms_string)); 
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   320
                           (warning ("initial thm: "^thmstring));
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   321
                           (warning ("subgoals: "^prems_string));
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   322
                           (warning ("pid: "^ pidstring))); 
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   323
                            isar_atp_aux thms thm (length prems) (childin, childout, pid) ;
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   324
                           
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   325
                           print_mode := (["xsymbols", "symbols"] @ ! print_mode)
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   326
    end;
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   327
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
   328
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   329
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   330
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   331
local
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   332
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   333
fun get_thms_cs claset =
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   334
    let val clsset = rep_cs claset
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   335
	val safeEs = #safeEs clsset
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   336
	val safeIs = #safeIs clsset
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   337
	val hazEs = #hazEs clsset
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   338
	val hazIs = #hazIs clsset
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   339
    in
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   340
	safeEs @ safeIs @ hazEs @ hazIs
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   341
    end;
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   342
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   343
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
   344
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   345
fun append_name name [] _ = []
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   346
  | 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
   347
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   348
fun append_names (name::names) (thms::thmss) =
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   349
    let val thms' = append_name name thms 0
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   350
    in
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   351
	thms'::(append_names names thmss)
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   352
    end;
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   353
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
   354
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   355
fun get_thms_ss [] = []
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   356
  | get_thms_ss thms =
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   357
    let val names = map Thm.name_of_thm thms 
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   358
        val thms' = map (mksimps mksimps_pairs) thms
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   359
        val thms'' = append_names names thms'
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   360
    in
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   361
	ResLib.flat_noDup thms''
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   362
    end;
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   363
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
   364
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   365
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   366
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   367
in
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   368
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
   369
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   370
(* convert locally declared rules to axiom clauses *)
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   371
(* write axiom clauses to ax_file *)
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   372
(* 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
   373
(* 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
   374
(*claset file and prob file*)
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   375
fun isar_local_thms (delta_cs, delta_ss_thms) =
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   376
    let val thms_cs = get_thms_cs delta_cs
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   377
	val thms_ss = get_thms_ss delta_ss_thms
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   378
	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
   379
	val clauses_strs = ResLib.flat_noDup (map ResClause.tptp_clause thms_clauses) (*string list*)
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   380
	val ax_file = File.sysify_path axiom_file
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   381
	val out = TextIO.openOut ax_file
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   382
    in
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   383
	(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
   384
    end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   385
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   386
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   387
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   388
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   389
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   390
(* called in Isar automatically *)
15679
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   391
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   392
fun isar_atp (ctxt,thm) =
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   393
    let val prems = ProofContext.prems_of ctxt
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   394
        val d_cs = Classical.get_delta_claset ctxt 
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   395
        val d_ss_thms = Simplifier.get_delta_simpset ctxt
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   396
        val thmstring = string_of_thm thm
15679
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   397
        val sg_prems = prems_of thm
15697
681bcb7f0389 removal of Main and other tidying up
paulson
parents: 15682
diff changeset
   398
        val sign = sign_of_thm thm
15679
28eb0fe50533 Integrating the reconstruction files into the building of HOL
quigley
parents: 15657
diff changeset
   399
        val prem_no = length sg_prems
15697
681bcb7f0389 removal of Main and other tidying up
paulson
parents: 15682
diff changeset
   400
        val prems_string =  concat_with_and (map (Sign.string_of_term sign) sg_prems) ""
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   401
    in
15644
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   402
         
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   403
          (warning ("initial thm in isar_atp: "^thmstring));
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   404
          (warning ("subgoals in isar_atp: "^prems_string));
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   405
    	   (warning ("number of subgoals in isar_atp: "^(string_of_int prem_no)));
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   406
          (isar_local_thms (d_cs,d_ss_thms); (warning("about to call isar_atp'"));
f2ef8c258fa4 *** empty log message ***
quigley
parents: 15608
diff changeset
   407
           isar_atp' (prems, thm))
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   408
    end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   409
15608
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   410
end
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   411
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   412
f161fa6f8fd5 bug fixes involving typechecking clauses
paulson
parents: 15603
diff changeset
   413
15452
e2a721567f67 Jia Meng: delta simpsets and clasets
paulson
parents: 15347
diff changeset
   414
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   415
end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   416
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   417
Proof.atp_hook := ResAtp.isar_atp;