diff -r 7a6f69cf5a86 -r a7a5157c5e75 src/HOL/Tools/ATP/recon_transfer_proof.ML --- a/src/HOL/Tools/ATP/recon_transfer_proof.ML Thu Sep 28 16:01:34 2006 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,638 +0,0 @@ -(* ID: $Id$ - Author: Claire Quigley - Copyright 2004 University of Cambridge -*) - -structure Recon_Transfer = -struct - -open Recon_Parse - -infixr 8 ++; infixr 7 >>; infixr 6 ||; - -val trace_path = Path.basic "transfer_trace"; - -fun trace s = if !Output.show_debug_msgs then File.append (File.tmp_path trace_path) s - else (); - - -(* Versions that include type information *) - -(* FIXME rename to str_of_thm *) -fun string_of_thm thm = - setmp show_sorts true (Pretty.str_of o Display.pretty_thm) thm; - - -(* check separate args in the watcher program for separating strings with a * or ; or something *) - -fun clause_strs_to_string [] str = str -| clause_strs_to_string (x::xs) str = clause_strs_to_string xs (str^x^"%") - -fun thmvars_to_string [] str = str -| thmvars_to_string (x::xs) str = thmvars_to_string xs (str^x^"%") - - -fun proofstep_to_string Axiom = "Axiom()" -| proofstep_to_string (Binary ((a,b), (c,d)))= - "Binary(("^(string_of_int a)^","^(string_of_int b)^"),("^(string_of_int c)^","^(string_of_int d)^"))" -| proofstep_to_string (Factor (a,b,c)) = - "Factor("^(string_of_int a)^","^(string_of_int b)^","^(string_of_int c)^")" -| proofstep_to_string (Para ((a,b), (c,d)))= - "Para(("^(string_of_int a)^","^(string_of_int b)^"),("^(string_of_int c)^","^(string_of_int d)^"))" -| proofstep_to_string (MRR ((a,b), (c,d))) = - "MRR(("^(string_of_int a)^","^(string_of_int b)^"),("^(string_of_int c)^","^(string_of_int d)^"))" -(*| proofstep_to_string (Rewrite((a,b),(c,d))) = - "Rewrite(("^(string_of_int a)^","^(string_of_int b)^"),("^(string_of_int c)^","^(string_of_int d)^"))"*) - - -fun proof_to_string (num,(step,clause_strs, thmvars)) = - (string_of_int num)^(proofstep_to_string step)^ - "["^(clause_strs_to_string clause_strs "")^"]["^(thmvars_to_string thmvars "")^"]" - - -fun proofs_to_string [] str = str -| proofs_to_string (x::xs) str = proofs_to_string xs (str ^ proof_to_string x) - - -fun init_proofstep_to_string (num, step, clause_strs) = - (string_of_int num)^" "^(proofstep_to_string step)^" "^ - (clause_strs_to_string clause_strs "")^" " - -fun init_proofsteps_to_string [] str = str -| init_proofsteps_to_string (x::xs) str = - init_proofsteps_to_string xs (str ^ init_proofstep_to_string x) - - - -(*** get a string representing the Isabelle ordered axioms ***) - -fun origAx_to_string (num,(meta,thmvars)) = - let val clause_strs = ReconOrderClauses.get_meta_lits_bracket meta - in - (string_of_int num)^"OrigAxiom()["^ - (clause_strs_to_string clause_strs "")^"]["^ - (thmvars_to_string thmvars "")^"]" - end - - -fun origAxs_to_string [] str = str -| origAxs_to_string (x::xs) str = - origAxs_to_string xs (str ^ origAx_to_string x ) - - -(*** get a string representing the Isabelle ordered axioms not used in the spass proof***) - -fun extraAx_to_string (num, (meta,thmvars)) = - let val clause_strs = ReconOrderClauses.get_meta_lits_bracket meta - in - (string_of_int num)^"ExtraAxiom()["^ - (clause_strs_to_string clause_strs "")^"]"^ - "["^(thmvars_to_string thmvars "")^"]" - end; - -fun extraAxs_to_string [] str = str -| extraAxs_to_string (x::xs) str = - let val newstr = extraAx_to_string x - in - extraAxs_to_string xs (str^newstr) - end; - -fun is_axiom (_,Axiom,str) = true -| is_axiom (_,_,_) = false - -fun get_step_nums [] nums = nums -| get_step_nums (( num:int,Axiom, str)::xs) nums = get_step_nums xs (nums@[num]) - -exception Noassoc; - -fun assoc_snd a [] = raise Noassoc - | assoc_snd a ((x, y)::t) = if a = y then x else assoc_snd a t; - -(* change to be something using check_order instead of a = y --> returns true if ASSERTION not raised in checkorder, false otherwise *) - -(*fun get_assoc_snds [] xs assocs= assocs -| get_assoc_snds (x::xs) ys assocs = get_assoc_snds xs ys (assocs@[((assoc_snd x ys))]) -*) -(*FIX - should this have vars in it? *) -fun there_out_of_order xs ys = (ReconOrderClauses.checkorder xs ys [] ([],[],[]); true) - handle _ => false - -fun assoc_out_of_order a [] = raise Noassoc -| assoc_out_of_order a ((b,c)::t) = if there_out_of_order a c then b else assoc_out_of_order a t; - -fun get_assoc_snds [] xs assocs= assocs -| get_assoc_snds (x::xs) ys assocs = get_assoc_snds xs ys (assocs@[((assoc_out_of_order x ys))]) - -fun add_if_not_inlist eq [] xs newlist = newlist -| add_if_not_inlist eq (y::ys) xs newlist = - if not (member eq xs y) then add_if_not_inlist eq ys xs (y::newlist) - else add_if_not_inlist eq ys xs newlist - -(*Flattens a list of list of strings to one string*) -fun onestr ls = String.concat (map String.concat ls); - - -(**** retrieve the axioms that were obtained from the clasimpset ****) - -(*Get names of axioms used. Axioms are indexed from 1, while the array is indexed from 0*) -fun get_axiom_names (names_arr: string array) step_nums = - let fun is_axiom n = n <= Array.length names_arr - fun index i = Array.sub(names_arr, i-1) - val axnums = List.filter is_axiom step_nums - val axnames = sort_distinct string_ord (map index axnums) - in - if length axnums = length step_nums then "UNSOUND!!" :: axnames - else axnames - end - - (*String contains multiple lines. We want those of the form - "253[0:Inp] et cetera..." - A list consisting of the first number in each line is returned. *) -fun get_spass_linenums proofstr = - let val toks = String.tokens (not o Char.isAlphaNum) - fun inputno (ntok::"0"::"Inp"::_) = Int.fromString ntok - | inputno _ = NONE - val lines = String.tokens (fn c => c = #"\n") proofstr - in List.mapPartial (inputno o toks) lines end - -fun get_axiom_names_spass proofstr names_arr = - get_axiom_names names_arr (get_spass_linenums proofstr); - - (*String contains multiple lines. We want those of the form - "number : ...: ...: initial..." *) -fun get_e_linenums proofstr = - let val fields = String.fields (fn c => c = #":") - val nospaces = String.translate (fn c => if c = #" " then "" else str c) - fun initial s = String.isPrefix "initial" (nospaces s) - fun firstno (line :: _ :: _ :: rule :: _) = - if initial rule then Int.fromString line else NONE - | firstno _ = NONE - val lines = String.tokens (fn c => c = #"\n") proofstr - in List.mapPartial (firstno o fields) lines end - -fun get_axiom_names_e proofstr names_arr = - get_axiom_names names_arr (get_e_linenums proofstr); - - (*String contains multiple lines. We want those of the form - "*********** [448, input] ***********". - A list consisting of the first number in each line is returned. *) -fun get_vamp_linenums proofstr = - let val toks = String.tokens (not o Char.isAlphaNum) - fun inputno [ntok,"input"] = Int.fromString ntok - | inputno _ = NONE - val lines = String.tokens (fn c => c = #"\n") proofstr - in List.mapPartial (inputno o toks) lines end - -fun get_axiom_names_vamp proofstr names_arr = - get_axiom_names names_arr (get_vamp_linenums proofstr); - - -(***********************************************) -(* get axioms for reconstruction *) -(***********************************************) -fun numclstr (vars, []) str = str -| numclstr ( vars, ((num, thm)::rest)) str = - let val newstr = str^(string_of_int num)^" "^(string_of_thm thm)^" " - in - numclstr (vars,rest) newstr - end - -fun addvars c (a,b) = (a,b,c) - -fun get_axioms_used proof_steps thms names_arr = - let val axioms = (List.filter is_axiom) proof_steps - val step_nums = get_step_nums axioms [] - - val clauses = make_clauses thms (*FIXME: must this be repeated??*) - - val vars = map thm_varnames clauses - - val distvars = distinct (op =) (fold append vars []) - val clause_terms = map prop_of clauses - val clause_frees = List.concat (map term_frees clause_terms) - - val frees = map lit_string_with_nums clause_frees; - - val distfrees = distinct (op =) frees - - val metas = map Meson.make_meta_clause clauses - val ax_strs = map #3 axioms - - (* literals of -all- axioms, not just those used by spass *) - val meta_strs = map ReconOrderClauses.get_meta_lits metas - - val metas_and_strs = ListPair.zip (metas,meta_strs) - val _ = trace ("\nAxioms: " ^ onestr ax_strs) - val _ = trace ("\nMeta_strs: " ^ onestr meta_strs) - - (* get list of axioms as thms with their variables *) - - val ax_metas = get_assoc_snds ax_strs metas_and_strs [] - val ax_vars = map thm_varnames ax_metas - val ax_with_vars = ListPair.zip (ax_metas,ax_vars) - - (* get list of extra axioms as thms with their variables *) - val extra_metas = add_if_not_inlist Thm.eq_thm metas ax_metas [] - val extra_vars = map thm_varnames extra_metas - val extra_with_vars = if not (null extra_metas) - then ListPair.zip (extra_metas,extra_vars) - else [] - in - (distfrees,distvars, extra_with_vars,ax_with_vars, ListPair.zip (step_nums,ax_metas)) - end; - - -(*********************************************************************) -(* Pass in spass string of proof and string version of isabelle goal *) -(* Get out reconstruction steps as a string to be sent to Isabelle *) -(*********************************************************************) - -fun rules_to_string [] = "NONE" - | rules_to_string xs = space_implode " " xs - - -(*The signal handler in watcher.ML must be able to read the output of this.*) -fun prover_lemma_list_aux getax proofstr probfile toParent ppid names_arr = - let val _ = trace - ("\n\nGetting lemma names. proofstr is " ^ proofstr ^ - "\nprobfile is " ^ probfile ^ - " num of clauses is " ^ string_of_int (Array.length names_arr)) - val axiom_names = getax proofstr names_arr - val ax_str = rules_to_string axiom_names - in - trace ("\nDone. Lemma list is " ^ ax_str); - TextIO.output (toParent, "Success. Lemmas used in automatic proof: " ^ - ax_str ^ "\n"); - TextIO.output (toParent, probfile ^ "\n"); - TextIO.flushOut toParent; - Posix.Process.kill(Posix.Process.K_PROC ppid, Posix.Signal.usr2) - end - handle exn => (*FIXME: exn handler is too general!*) - (trace ("\nprover_lemma_list_aux: In exception handler: " ^ - Toplevel.exn_message exn); - TextIO.output (toParent, "Translation failed for the proof: " ^ - String.toString proofstr ^ "\n"); - TextIO.output (toParent, probfile); - TextIO.flushOut toParent; - Posix.Process.kill(Posix.Process.K_PROC ppid, Posix.Signal.usr2)); - -val e_lemma_list = prover_lemma_list_aux get_axiom_names_e; - -val vamp_lemma_list = prover_lemma_list_aux get_axiom_names_vamp; - -val spass_lemma_list = prover_lemma_list_aux get_axiom_names_spass; - - -(**** Full proof reconstruction for SPASS (not really working) ****) - -fun spass_reconstruct proofstr probfile toParent ppid thms names_arr = - let val _ = trace ("\nspass_reconstruct. Proofstr is "^proofstr) - val tokens = #1(lex proofstr) - - (* parse spass proof into datatype *) - (***********************************) - val proof_steps = parse tokens - val _ = trace "\nParsing finished" - - (************************************) - (* recreate original subgoal as thm *) - (************************************) - (* get axioms as correctly numbered clauses w.r.t. the Spass proof *) - (* need to get prems_of thm, then get right one of the prems, relating to whichever*) - (* subgoal this is, and turn it into meta_clauses *) - (* should prob add array and table here, so that we can get axioms*) - (* produced from the clasimpset rather than the problem *) - val (frees,vars,extra_with_vars ,ax_with_vars,numcls) = get_axioms_used proof_steps thms names_arr - - (*val numcls_string = numclstr ( vars, numcls) ""*) - val _ = trace "\ngot axioms" - - (************************************) - (* translate proof *) - (************************************) - val _ = trace ("\nabout to translate proof, steps: " - ^ (init_proofsteps_to_string proof_steps "")) - val (newthm,proof) = translate_proof numcls proof_steps vars - val _ = trace ("translated proof, steps: "^(init_proofsteps_to_string proof_steps "")) - (***************************************************) - (* transfer necessary steps as strings to Isabelle *) - (***************************************************) - (* turn the proof into a string *) - val reconProofStr = proofs_to_string proof "" - (* do the bit for the Isabelle ordered axioms at the top *) - val ax_nums = map #1 numcls - val ax_strs = map ReconOrderClauses.get_meta_lits_bracket (map #2 numcls) - val numcls_strs = ListPair.zip (ax_nums,ax_strs) - val num_cls_vars = map (addvars vars) numcls_strs; - val reconIsaAxStr = origAxs_to_string (ListPair.zip (ax_nums,ax_with_vars)) "" - - val extra_nums = if not (null extra_with_vars) then (1 upto (length extra_with_vars)) - else [] - val reconExtraAxStr = extraAxs_to_string ( ListPair.zip (extra_nums,extra_with_vars)) "" - val frees_str = "["^(thmvars_to_string frees "")^"]" - val reconstr = (frees_str^reconExtraAxStr^reconIsaAxStr^reconProofStr) - val _ = trace ("\nReconstruction:\n" ^ reconstr) - in - TextIO.output (toParent, reconstr^"\n"); - TextIO.output (toParent, probfile ^ "\n"); - TextIO.flushOut toParent; - Posix.Process.kill(Posix.Process.K_PROC ppid, Posix.Signal.usr2); - all_tac - end - handle exn => (*FIXME: exn handler is too general!*) - (trace ("\nspass_reconstruct. In exception handler: " ^ Toplevel.exn_message exn); - TextIO.output (toParent,"Translation failed for SPASS proof:"^ - String.toString proofstr ^"\n"); - TextIO.output (toParent, probfile ^ "\n"); - TextIO.flushOut toParent; - Posix.Process.kill(Posix.Process.K_PROC ppid, Posix.Signal.usr2); all_tac) - -(**********************************************************************************) -(* At other end, want to turn back into datatype so can apply reconstruct_proof. *) -(* This will be done by the signal handler *) -(**********************************************************************************) - -(* Parse in the string version of the proof steps for reconstruction *) -(* Isar format: cl1 [BINARY 0 cl2 0];cl1 [PARAMOD 0 cl2 0]; cl1 [DEMOD 0 cl2];cl1 [FACTOR 1 2];*) - - - val term_numstep = - (number ++ (a (Other ",")) ++ number) >> (fn (a, (_, c)) => (a, c)) - -val extraaxiomstep = (a (Word "ExtraAxiom"))++ (a (Other "(")) ++(a (Other ")")) - >> (fn (_) => ExtraAxiom) - - - -val origaxiomstep = (a (Word "OrigAxiom"))++ (a (Other "(")) ++(a (Other ")")) - >> (fn (_) => OrigAxiom) - - - val axiomstep = (a (Word "Axiom"))++ (a (Other "(")) ++(a (Other ")")) - >> (fn (_) => Axiom) - - - - - val binarystep = (a (Word "Binary")) ++ (a (Other "(")) ++ (a (Other "(")) - ++ term_numstep ++ (a (Other ")")) ++ (a (Other ",")) - ++ (a (Other "(")) ++ term_numstep ++ (a (Other ")")) ++ (a (Other ")")) - >> (fn (_, (_, (_, (c, (_,(_,(_, (e,(_,_))))))))) => Binary (c,e)) - - - val parastep = (a (Word "Para")) ++ (a (Other "(")) ++ (a (Other "(")) - ++ term_numstep ++ (a (Other ")")) ++ (a (Other ",")) - ++ (a (Other "(")) ++ term_numstep ++ (a (Other ")")) ++ (a (Other ")")) - >> (fn (_, (_, (_, (c, (_,(_,(_, (e,(_,_))))))))) => Para(c, e)) - - val mrrstep = (a (Word "MRR")) ++ (a (Other "(")) ++ (a (Other "(")) - ++ term_numstep ++ (a (Other ")")) ++ (a (Other ",")) - ++ (a (Other "(")) ++ term_numstep ++ (a (Other ")")) ++ (a (Other ")")) - >> (fn (_, (_, (_, (c, (_,(_,(_, (e,(_,_))))))))) => MRR(c, e)) - - - val factorstep = (a (Word "Factor")) ++ (a (Other "(")) - ++ number ++ (a (Other ",")) - ++ number ++ (a (Other ",")) - ++ number ++ (a (Other ")")) - - >> (fn (_, (_, (c, (_, (e,(_,(f,_))))))) => Factor (c,e,f)) - - -(*val rewritestep = (a (Word "Rewrite")) ++ (a (Other "(")) ++ (a (Other "(")) - ++ term_numstep ++ (a (Other ")")) ++ (a (Other ",")) - ++ (a (Other "(")) ++ term_numstep ++ (a (Other ")")) ++ (a (Other ")")) - >> (fn (_, (_, (_, (c, (_,(_,(_, (e,(_,_))))))))) => Rewrite (c,e))*) - -val obviousstep = (a (Word "Obvious")) ++ (a (Other "(")) - ++ term_numstep ++ (a (Other ")")) - >> (fn (_, (_, (c,_))) => Obvious (c)) - - val methodstep = extraaxiomstep || origaxiomstep || axiomstep ||binarystep || factorstep|| parastep || mrrstep || (*rewritestep ||*) obviousstep - - - val number_list_step = - ( number ++ many ((a (Other ",") ++ number)>> #2)) - >> (fn (a,b) => (a::b)) - - val numberlist_step = a (Other "[") ++ a (Other "]") - >>(fn (_,_) => ([]:int list)) - || a (Other "[") ++ number_list_step ++ a (Other "]") - >>(fn (_,(a,_)) => a) - - - -(** change this to allow P (x U) *) - fun arglist_step input = - ( word ++ many word >> (fn (a, b) => (a^" "^(space_implode " " b))) - ||word >> (fn (a) => (a)))input - - -fun literal_step input = (word ++ a (Other "(") ++ arglist_step ++ a (Other ")") - >>(fn (a, (b, (c,d))) => (a^" ("^(c)^")")) - || arglist_step >> (fn (a) => (a)))input - - - -(* fun term_step input = (a (Other "~") ++ arglist_step ++ a (Other "%")>> (fn (a,(b,c)) => ("~ "^b)) - || arglist_step ++ a (Other "%")>> (fn (a,b) => a ))input -*) - - - fun term_step input = (a (Other "~") ++ literal_step ++ a (Other "%")>> (fn (a,(b,c)) => ("~ "^b)) - || literal_step ++ a (Other "%")>> (fn (a,b) => a ))input - - - - - val term_list_step = - ( term_step ++ many ( term_step)) - >> (fn (a,b) => (a::b)) - - -val term_lists_step = a (Other "[") ++ a (Other "]") - >>(fn (_,_) => ([]:string list)) - || a (Other "[") ++ term_list_step ++ a (Other "]") - >>(fn (_,(a,_)) => a) - - - val linestep = number ++ methodstep ++ term_lists_step ++ term_lists_step - >> (fn (a, (b, (c,d))) => (a,(b,c,d))) - - val lines_step = many linestep - - val alllines_step = (term_lists_step ++ lines_step ) ++ finished >> #1 - - val parse_step = #1 o alllines_step - - - (* -val reconstr ="[P%x%xa%xb%]1OrigAxiom()[P x%~ P U%][U%]3OrigAxiom()[P U%~ P x%][U%]5OrigAxiom()[~ P xa%~ P U%][U%]7OrigAxiom()[P U%P xb%][U%]1Axiom()[P x%~ P U%][U%]3Axiom()[P U%~ P x%][U%]5Axiom()[~ P U%~ P xa%][U%]7Axiom()[P U%P xb%][U%]9Factor(5,0,1)[~ P xa%][]10Binary((9,0),(3,0))[~ P x%][]11Binary((10,0),(1,0))[~ P U%][U%]12Factor(7,0,1)[P xb%][]14Binary((11,0),(12,0))[][]%(EX x::'a::type. ALL y::'a::type. (P::'a::type => bool) x = P y) -->(EX x::'a::type. P x) = (ALL y::'a::type. P y)" -*) - -(************************************************************) -(* Construct an Isar style proof from a list of proof steps *) -(************************************************************) -(* want to assume all axioms, then do haves for the other clauses*) -(* then show for the last step *) - -(* replace ~ by not here *) -val change_nots = String.translate (fn c => if c = #"~" then "\\" else str c); - -fun clstrs_to_string xs = space_implode "; " (map change_nots xs); - -fun thmvars_to_quantstring [] str = str -| thmvars_to_quantstring (x::[]) str =str^x^". " -| thmvars_to_quantstring (x::xs) str = thmvars_to_quantstring xs (str^(x^" ")) - - -fun clause_strs_to_isar clstrs [] = - "\"\\"^(clstrs_to_string clstrs)^"\\ \\ False\"" -| clause_strs_to_isar clstrs thmvars = - "\"\\"^(thmvars_to_quantstring thmvars "")^ - "\\"^(clstrs_to_string clstrs)^"\\ \\ False\"" - -fun frees_to_isar_str clstrs = space_implode " " (map change_nots clstrs) - - -(***********************************************************************) -(* functions for producing assumptions for the Isabelle ordered axioms *) -(***********************************************************************) -(*val str = "[P%x%xa%xb%]1OrigAxiom()[P x%~ P U%][U%]3OrigAxiom()[P U%~ P x%][U%]5OrigAxiom()[~ P xa%~ P U%][U%]7OrigAxiom()[P U%P xb%][U%]1Axiom()[P x%~ P U%][U%]3Axiom()[P U%~ P x%][U%]5Axiom()[~ P U%~ P xa%][U%]7Axiom()[P U%P xb%][U%]9Factor(5,0,1)[~ P xa%][]10Binary((9,0),(3,0))[~ P x%][]11Binary((10,0),(1,0))[~ P U%][U%]12Factor(7,0,1)[P xb%][]14Binary((11,0),(12,0))[][]"; -num, rule, clausestrs, vars*) - - -(* assume the extra clauses - not used in Spass proof *) - -fun is_extraaxiom_step ( num:int,(ExtraAxiom, str, tstr)) = true -| is_extraaxiom_step (num, _) = false - -fun get_extraaxioms xs = List.filter (is_extraaxiom_step) ( xs) - -fun assume_isar_extraaxiom [] str = str -| assume_isar_extraaxiom ((numb,(step, clstr, thmvars))::xs) str = assume_isar_extraaxiom xs (str^"and cl"^(string_of_int numb)^"': "^(clause_strs_to_isar clstr thmvars)^"\n " ) - - - -fun assume_isar_extraaxioms [] = "" -|assume_isar_extraaxioms ((numb,(step, clstrs, thmstrs))::xs) = let val str = "assume cl"^(string_of_int numb)^"': "^(clause_strs_to_isar clstrs thmstrs)^"\n" - in - assume_isar_extraaxiom xs str - end - -(* assume the Isabelle ordered clauses *) - -fun is_origaxiom_step ( num:int,(OrigAxiom, str, tstr)) = true -| is_origaxiom_step (num, _) = false - -fun get_origaxioms xs = List.filter (is_origaxiom_step) ( xs) - -fun assume_isar_origaxiom [] str = str -| assume_isar_origaxiom ((numb,(step, clstr, thmvars))::xs) str = assume_isar_origaxiom xs (str^"and cl"^(string_of_int numb)^"': "^(clause_strs_to_isar clstr thmvars)^"\n " ) - - - -fun assume_isar_origaxioms ((numb,(step, clstrs, thmstrs))::xs) = let val str = "assume cl"^(string_of_int numb)^"': "^(clause_strs_to_isar clstrs thmstrs)^"\n" - in - assume_isar_origaxiom xs str - end - - - -fun is_axiom_step ( num:int,(Axiom, str, tstr)) = true -| is_axiom_step (num, _) = false - -fun get_axioms xs = List.filter (is_axiom_step) ( xs) - -fun have_isar_axiomline (numb,(step, clstrs, thmstrs))="have cl"^(string_of_int numb)^": "^(clause_strs_to_isar clstrs thmstrs)^"\n" - -fun by_isar_axiomline (numb,(step, clstrs, thmstrs))="by (rule cl"^ (string_of_int numb)^"') \n" - - -fun isar_axiomline (numb, (step, clstrs, thmstrs)) = (have_isar_axiomline (numb,(step,clstrs, thmstrs )))^( by_isar_axiomline(numb,(step,clstrs, thmstrs )) ) - - -fun isar_axiomlines [] str = str -| isar_axiomlines (x::xs) str = isar_axiomlines xs (str^(isar_axiomline x)) - - -fun have_isar_line (numb,(step, clstrs, thmstrs))="have cl"^(string_of_int numb)^": "^(clause_strs_to_isar clstrs thmstrs)^"\n" -(*FIX: ask Larry to add and mrr attribute *) - -fun by_isar_line ((Binary ((a,b), (c,d)))) = - "by(rule cl"^ - (string_of_int a)^" [binary "^(string_of_int b)^" cl"^ - (string_of_int c)^" "^(string_of_int d)^"])\n" -|by_isar_line ((MRR ((a,b), (c,d)))) = - "by(rule cl"^ - (string_of_int a)^" [binary "^(string_of_int b)^" cl"^ - (string_of_int c)^" "^(string_of_int d)^"])\n" -| by_isar_line ( (Para ((a,b), (c,d)))) = - "by (rule cl"^ - (string_of_int a)^" [paramod "^(string_of_int b)^" cl"^ - (string_of_int c)^" "^(string_of_int d)^"])\n" -| by_isar_line ((Factor ((a,b,c)))) = - "by (rule cl"^(string_of_int a)^" [factor "^(string_of_int b)^" "^ - (string_of_int c)^" ])\n" -(*| by_isar_line ( (Rewrite ((a,b),(c,d)))) = - "by (rule cl"^(string_of_int a)^" [demod "^(string_of_int b)^" "^ - (string_of_int c)^" "^(string_of_int d)^" ])\n"*) -| by_isar_line ( (Obvious ((a,b)))) = - "by (rule cl"^(string_of_int a)^" [obvious "^(string_of_int b)^" ])\n" - -fun isar_line (numb, (step, clstrs, thmstrs)) = (have_isar_line (numb,(step,clstrs, thmstrs )))^( by_isar_line step) - - -fun isar_lines [] str = str -| isar_lines (x::xs) str = isar_lines xs (str^(isar_line x)) - -fun last_isar_line (numb,( step, clstrs,thmstrs)) = "show \"False\"\n"^(by_isar_line step) - - -fun to_isar_proof (frees, xs) = - let val extraaxioms = get_extraaxioms xs - val extraax_num = length extraaxioms - val origaxioms_and_steps = Library.drop (extraax_num, xs) - - val origaxioms = get_origaxioms origaxioms_and_steps - val origax_num = length origaxioms - val axioms_and_steps = Library.drop (origax_num + extraax_num, xs) - val axioms = get_axioms axioms_and_steps - - val steps = Library.drop (origax_num, axioms_and_steps) - val (firststeps, laststep) = split_last steps - - val isar_proof = - ("show \"[your goal]\"\n")^ - ("proof (rule ccontr,skolemize, make_clauses) \n")^ - ("fix "^(frees_to_isar_str frees)^"\n")^ - (assume_isar_extraaxioms extraaxioms)^ - (assume_isar_origaxioms origaxioms)^ - (isar_axiomlines axioms "")^ - (isar_lines firststeps "")^ - (last_isar_line laststep)^ - ("qed") - val _ = trace ("\nto_isar_proof returns " ^ isar_proof) - in - isar_proof - end; - -(* get fix vars from axioms - all Frees *) -(* check each clause for meta-vars and /\ over them at each step*) - -(*******************************************************) -(* This assumes the thm list "numcls" is still there *) -(* In reality, should probably label it with an *) -(* ID number identifying the subgoal. This could *) -(* be passed over to the watcher, e.g. numcls25 *) -(*******************************************************) - -fun apply_res_thm str = - let val tokens = #1 (lex str); - val _ = trace ("\napply_res_thm. str is: "^str^"\n") - val (frees,recon_steps) = parse_step tokens - in - to_isar_proof (frees, recon_steps) - end - -end;