src/HOL/Tools/res_atp.ML
changeset 17150 ce2a1aeb42aa
parent 17091 13593aa6a546
child 17231 f42bc4f7afdf
equal deleted inserted replaced
17149:e2b19c92ef51 17150:ce2a1aeb42aa
    14 (*val atp_tac : int -> Tactical.tactic*)
    14 (*val atp_tac : int -> Tactical.tactic*)
    15   val full_spass: bool ref
    15   val full_spass: bool ref
    16 (*val spass: bool ref*)
    16 (*val spass: bool ref*)
    17   val vampire: bool ref
    17   val vampire: bool ref
    18   val custom_spass: string list ref
    18   val custom_spass: string list ref
       
    19   val hook_count: int ref
       
    20 (*  val invoke_atp: Toplevel.transition -> Toplevel.transition*)
    19 end;
    21 end;
    20 
    22 
    21 structure ResAtp: RES_ATP =
    23 structure ResAtp: RES_ATP =
    22 struct
    24 struct
       
    25 
       
    26 
       
    27 val call_atp = ref false;
       
    28 val hook_count = ref 0;
    23 
    29 
    24 fun debug_tac tac = (debug "testing"; tac);
    30 fun debug_tac tac = (debug "testing"; tac);
    25 
    31 
    26 val full_spass = ref false;
    32 val full_spass = ref false;
    27 
    33 
    54 val axiom_file = File.tmp_path (Path.basic "axioms");
    60 val axiom_file = File.tmp_path (Path.basic "axioms");
    55 val clasimp_file = File.tmp_path (Path.basic "clasimp");
    61 val clasimp_file = File.tmp_path (Path.basic "clasimp");
    56 val hyps_file = File.tmp_path (Path.basic "hyps");
    62 val hyps_file = File.tmp_path (Path.basic "hyps");
    57 val prob_file = File.tmp_path (Path.basic "prob");
    63 val prob_file = File.tmp_path (Path.basic "prob");
    58 val dummy_tac = all_tac;
    64 val dummy_tac = all_tac;
       
    65 val _ =debug  (File.platform_path prob_file);
    59 
    66 
    60 
    67 
    61 (**** for Isabelle/ML interface  ****)
    68 (**** for Isabelle/ML interface  ****)
    62 
    69 
    63 (*Remove unwanted characters such as ? and newline from the textural 
    70 (*Remove unwanted characters such as ? and newline from the textural 
   133 
   140 
   134 (*********************************************************************)
   141 (*********************************************************************)
   135 (* write out a subgoal as DFG clauses to the file "probN"           *)
   142 (* write out a subgoal as DFG clauses to the file "probN"           *)
   136 (* where N is the number of this subgoal                             *)
   143 (* where N is the number of this subgoal                             *)
   137 (*********************************************************************)
   144 (*********************************************************************)
   138 (*
   145 
   139 fun dfg_inputs_tfrees thms n tfrees = 
   146 (*fun dfg_inputs_tfrees thms n tfrees axclauses= 
   140     let val _ = (debug ("in dfg_inputs_tfrees 0"))
   147     let  val clss = map (ResClause.make_conjecture_clause_thm) thms
   141         val clss = map (ResClause.make_conjecture_clause_thm) thms
       
   142          val _ = (debug ("in dfg_inputs_tfrees 1"))
   148          val _ = (debug ("in dfg_inputs_tfrees 1"))
   143 	val (dfg_clss,tfree_litss) = ListPair.unzip (map ResClause.clause2dfg clss)
   149 	 val (dfg_clss,tfree_litss) = ListPair.unzip (map ResClause.clause2dfg clss)
   144         val _ = (debug ("in dfg_inputs_tfrees 2"))
   150 	 val tfree_clss = map ResClause.tfree_clause ((ResLib.flat_noDup tfree_litss) \\ tfrees) 
   145 	val tfree_clss = map ResClause.tfree_clause ((ResLib.flat_noDup tfree_litss) \\ tfrees) 
   151          val probfile = (File.platform_path prob_file) ^ "_" ^ (string_of_int n)
   146          val _ = (debug ("in dfg_inputs_tfrees 3"))
   152          val filestr = ResClause.clauses2dfg clss ("prob" ^ (string_of_int n)) axclauses [] [] [] tfrees   
       
   153 	 val out = TextIO.openOut(probfile)
       
   154     in
       
   155 	(ResLib.writeln_strs out (tfree_clss @ dfg_clss); TextIO.closeOut out; debug probfile)
       
   156     end;
       
   157 *)
       
   158 fun dfg_inputs_tfrees thms n tfrees axclauses = 
       
   159     let val clss = map (ResClause.make_conjecture_clause_thm) thms
   147         val probfile = (File.platform_path prob_file) ^ "_" ^ (string_of_int n)
   160         val probfile = (File.platform_path prob_file) ^ "_" ^ (string_of_int n)
       
   161         val _ = warning ("about to write out dfg prob file "^probfile)
       
   162        	(*val (dfg_clss,tfree_litss) = ListPair.unzip (map ResClause.clause2dfg clss)
       
   163         val tfree_clss = map ResClause.tfree_dfg_clause ((ResLib.flat_noDup tfree_litss) \\ tfrees) *)   
       
   164         val filestr = ResClause.clauses2dfg clss ("prob" ^ (string_of_int n)) axclauses [] [] [] tfrees   
   148 	val out = TextIO.openOut(probfile)
   165 	val out = TextIO.openOut(probfile)
   149     in
   166     in
   150 	(ResLib.writeln_strs out (tfree_clss @ dfg_clss); TextIO.closeOut out; debug probfile
   167 	(ResLib.writeln_strs out [filestr]; TextIO.closeOut out; debug probfile )
   151     end;*)
   168     end;
       
   169 
   152 
   170 
   153 (*********************************************************************)
   171 (*********************************************************************)
   154 (* call SPASS with settings and problem file for the current subgoal *)
   172 (* call SPASS with settings and problem file for the current subgoal *)
   155 (* should be modified to allow other provers to be called            *)
   173 (* should be modified to allow other provers to be called            *)
   156 (*********************************************************************)
   174 (*********************************************************************)
   208 (**********************************************************)
   226 (**********************************************************)
   209 (* write out the current subgoal as a tptp file, probN,   *)
   227 (* write out the current subgoal as a tptp file, probN,   *)
   210 (* then call dummy_tac - should be call_res_tac           *)
   228 (* then call dummy_tac - should be call_res_tac           *)
   211 (**********************************************************)
   229 (**********************************************************)
   212 
   230 
   213 fun get_sko_thms tfrees sign sg_terms (childin, childout, pid) thm n sko_thms =
   231 
   214   if n = 0 then
   232 fun get_sko_thms tfrees sign sg_terms (childin, childout,pid) thm n sko_thms axclauses =
   215     (call_resolve_tac (rev sko_thms)
   233     if n=0 then 
   216       sign sg_terms (childin, childout, pid) (List.length sg_terms);
   234        (call_resolve_tac  (rev sko_thms)
   217      dummy_tac thm)
   235         sign  sg_terms (childin, childout, pid) (List.length sg_terms);
   218   else
   236         dummy_tac thm)
   219     SELECT_GOAL
   237      else
   220       (EVERY1 [rtac ccontr, atomize_tac, skolemize_tac,
   238 	
   221         METAHYPS (fn negs =>
   239      ( SELECT_GOAL
   222           (tptp_inputs_tfrees (make_clauses negs) n tfrees;
   240         (EVERY1 [rtac ccontr,atomize_tac, skolemize_tac, 
   223            get_sko_thms tfrees sign sg_terms (childin, childout, pid) thm (n - 1)
   241           METAHYPS(fn negs => (if !SpassComm.spass then 
   224              (negs::sko_thms); dummy_tac))]) n thm;
   242 				    dfg_inputs_tfrees (make_clauses negs) n tfrees axclauses
       
   243  				 else
       
   244 			  	    tptp_inputs_tfrees (make_clauses negs)  n tfrees;
       
   245 			            get_sko_thms  tfrees sign sg_terms (childin, childout, pid) thm  (n -1) (negs::sko_thms) axclauses; dummy_tac))]) n thm )
       
   246 
   225 
   247 
   226 
   248 
   227 (**********************************************)
   249 (**********************************************)
   228 (* recursively call atp_tac_g on all subgoals *)
   250 (* recursively call atp_tac_g on all subgoals *)
   229 (* sg_term is the nth subgoal as a term - used*)
   251 (* sg_term is the nth subgoal as a term - used*)
   230 (* in proof reconstruction                    *)
   252 (* in proof reconstruction                    *)
   231 (**********************************************)
   253 (**********************************************)
   232 
   254 
   233 fun isar_atp_goal' thm n tfree_lits (childin, childout, pid) =
   255 fun isar_atp_goal' thm n tfree_lits (childin, childout, pid)  axclauses =
   234   let
   256   let
   235     val prems = Thm.prems_of thm
   257     val prems = Thm.prems_of thm
   236     (*val sg_term = get_nth k prems*)
   258     (*val sg_term = get_nth k prems*)
   237     val sign = sign_of_thm thm
   259     val sign = sign_of_thm thm
   238     val thmstring = string_of_thm thm
   260     val thmstring = string_of_thm thm
   241     debug("thmstring in isar_atp_goal': " ^ thmstring);
   263     debug("thmstring in isar_atp_goal': " ^ thmstring);
   242     (* go and call callResProvers with this subgoal *)
   264     (* go and call callResProvers with this subgoal *)
   243     (* isar_atp_g tfree_lits  sg_term (childin, childout, pid) k thm; *)
   265     (* isar_atp_g tfree_lits  sg_term (childin, childout, pid) k thm; *)
   244     (* recursive call to pick up the remaining subgoals *)
   266     (* recursive call to pick up the remaining subgoals *)
   245     (* isar_atp_goal' thm (k+1) n tfree_lits  (childin, childout, pid) *)
   267     (* isar_atp_goal' thm (k+1) n tfree_lits  (childin, childout, pid) *)
   246     get_sko_thms tfree_lits sign prems (childin, childout, pid) thm n []
   268     get_sko_thms tfree_lits sign prems (childin, childout, pid) thm n []  axclauses
   247   end;
   269   end;
   248 
   270 
   249 
   271 
   250 (**************************************************)
   272 (**************************************************)
   251 (* convert clauses from "assume" to conjecture.   *)
   273 (* convert clauses from "assume" to conjecture.   *)
   253 (* any hypotheses in the goal produced by assume  *)
   275 (* any hypotheses in the goal produced by assume  *)
   254 (* statements;                                    *)
   276 (* statements;                                    *)
   255 (* write to file "hyps"                           *)
   277 (* write to file "hyps"                           *)
   256 (**************************************************)
   278 (**************************************************)
   257 
   279 
   258 fun isar_atp_aux thms thm n_subgoals  (childin, childout, pid) =
   280 fun isar_atp_aux thms thm n_subgoals  (childin, childout, pid)  axclauses=
   259   let val tfree_lits = isar_atp_h thms
   281   let val tfree_lits = isar_atp_h thms
   260   in
   282   in
   261     debug ("in isar_atp_aux");
   283     debug ("in isar_atp_aux");
   262     isar_atp_goal' thm n_subgoals tfree_lits (childin, childout, pid)
   284     isar_atp_goal' thm n_subgoals tfree_lits (childin, childout, pid)  axclauses
   263   end;
   285   end;
   264 
   286 
   265 (******************************************************************)
   287 (******************************************************************)
   266 (* called in Isar automatically                                   *)
   288 (* called in Isar automatically                                   *)
   267 (* writes out the current clasimpset to a tptp file               *)
   289 (* writes out the current clasimpset to a tptp file               *)
   279       val thms_string = Meson.concat_with_and (map string_of_thm thms)
   301       val thms_string = Meson.concat_with_and (map string_of_thm thms)
   280       val thm_string = string_of_thm thm
   302       val thm_string = string_of_thm thm
   281       val prems_string = Meson.concat_with_and (map (Sign.string_of_term thy) prems)
   303       val prems_string = Meson.concat_with_and (map (Sign.string_of_term thy) prems)
   282 
   304 
   283       (*set up variables for writing out the clasimps to a tptp file*)
   305       (*set up variables for writing out the clasimps to a tptp file*)
   284       val (clause_arr, num_of_clauses) =
   306       val (clause_arr, num_of_clauses, axclauses) =
   285         ResClasimp.write_out_clasimp (File.platform_path clasimp_file) thy
   307         ResClasimp.write_out_clasimp (File.platform_path clasimp_file) thy
   286           (hd prems) (*FIXME: hack!! need to do all prems*)
   308           (hd prems) (*FIXME: hack!! need to do all prems*)
   287       val _ = debug ("clasimp_file is " ^ File.platform_path clasimp_file)
   309       val _ = debug ("clasimp_file is " ^ File.platform_path clasimp_file ^ " with " ^ (string_of_int num_of_clauses)^ " clauses")
   288       val (childin, childout, pid) = Watcher.createWatcher (thm, clause_arr, num_of_clauses)
   310       val (childin, childout, pid) = Watcher.createWatcher (thm, clause_arr, num_of_clauses)
   289       val pid_string =
   311       val pid_string =
   290         string_of_int (Word.toInt (Word.fromLargeWord (Posix.Process.pidToWord pid)))
   312         string_of_int (Word.toInt (Word.fromLargeWord (Posix.Process.pidToWord pid)))
   291     in
   313     in
   292       debug ("initial thms: " ^ thms_string);
   314       debug ("initial thms: " ^ thms_string);
   293       debug ("initial thm: " ^ thm_string);
   315       debug ("initial thm: " ^ thm_string);
   294       debug ("subgoals: " ^ prems_string);
   316       debug ("subgoals: " ^ prems_string);
   295       debug ("pid: "^ pid_string);
   317       debug ("pid: "^ pid_string);
   296       isar_atp_aux thms thm (length prems) (childin, childout, pid);
   318       isar_atp_aux thms thm (length prems) (childin, childout, pid) axclauses;
   297       ()
   319       ()
   298     end);
   320     end);
   299 
   321 
   300 
   322 
   301 fun get_thms_cs claset =
   323 fun get_thms_cs claset =
   338 
   360 
   339 
   361 
   340 (** the Isar toplevel hook **)
   362 (** the Isar toplevel hook **)
   341 
   363 
   342 val invoke_atp = Toplevel.unknown_proof o Toplevel.keep (fn state =>
   364 val invoke_atp = Toplevel.unknown_proof o Toplevel.keep (fn state =>
   343   let
   365 
       
   366   let
       
   367 
   344     val proof = Toplevel.proof_of state
   368     val proof = Toplevel.proof_of state
   345     val (ctxt, (_, goal)) = Proof.get_goal proof
   369     val (ctxt, (_, goal)) = Proof.get_goal proof
   346         handle Proof.STATE _ => error "No goal present";
   370         handle Proof.STATE _ => error "No goal present";
       
   371 
   347     val thy = ProofContext.theory_of ctxt;
   372     val thy = ProofContext.theory_of ctxt;
   348 
   373 
   349     (* FIXME presently unused *)
   374     (* FIXME presently unused *)
   350     val ss_thms = subtract_simpset thy ctxt;
   375     val ss_thms = subtract_simpset thy ctxt;
   351     val cs_thms = subtract_claset thy ctxt;
   376     val cs_thms = subtract_claset thy ctxt;
   354            Pretty.string_of (ProofContext.pretty_thm ctxt goal));
   379            Pretty.string_of (ProofContext.pretty_thm ctxt goal));
   355     debug ("subgoals in isar_atp: " ^ 
   380     debug ("subgoals in isar_atp: " ^ 
   356            Pretty.string_of (ProofContext.pretty_term ctxt
   381            Pretty.string_of (ProofContext.pretty_term ctxt
   357              (Logic.mk_conjunction_list (Thm.prems_of goal))));
   382              (Logic.mk_conjunction_list (Thm.prems_of goal))));
   358     debug ("number of subgoals in isar_atp: " ^ string_of_int (Thm.nprems_of goal));
   383     debug ("number of subgoals in isar_atp: " ^ string_of_int (Thm.nprems_of goal));
       
   384     hook_count := !hook_count +1;
       
   385     debug ("in hook for time: " ^(string_of_int (!hook_count)) );
   359     ResClause.init thy;
   386     ResClause.init thy;
   360     isar_atp' (ctxt, ProofContext.prems_of ctxt, goal)
   387     isar_atp' (ctxt, ProofContext.prems_of ctxt, goal)
   361   end);
   388   end);
   362 
   389 
   363 val call_atpP =
   390 val call_atpP =