src/HOL/Tools/ATP/recon_transfer_proof.ML
changeset 20762 a7a5157c5e75
parent 20761 7a6f69cf5a86
child 20763 052b348a98a9
equal deleted inserted replaced
20761:7a6f69cf5a86 20762:a7a5157c5e75
     1 (*  ID:         $Id$
       
     2     Author:     Claire Quigley
       
     3     Copyright   2004  University of Cambridge
       
     4 *)
       
     5 
       
     6 structure Recon_Transfer =
       
     7 struct
       
     8 
       
     9 open Recon_Parse
       
    10 
       
    11 infixr 8 ++; infixr 7 >>; infixr 6 ||;
       
    12 
       
    13 val trace_path = Path.basic "transfer_trace";
       
    14 
       
    15 fun trace s = if !Output.show_debug_msgs then File.append (File.tmp_path trace_path) s 
       
    16               else ();
       
    17 
       
    18 
       
    19 (* Versions that include type information *)
       
    20  
       
    21 (* FIXME rename to str_of_thm *)
       
    22 fun string_of_thm thm =
       
    23   setmp show_sorts true (Pretty.str_of o Display.pretty_thm) thm;
       
    24 
       
    25 
       
    26 (* check separate args in the watcher program for separating strings with a * or ; or something *)
       
    27 
       
    28 fun clause_strs_to_string [] str = str
       
    29 |   clause_strs_to_string (x::xs) str = clause_strs_to_string xs (str^x^"%")
       
    30 
       
    31 fun thmvars_to_string [] str = str
       
    32 |   thmvars_to_string (x::xs) str = thmvars_to_string xs (str^x^"%")
       
    33 
       
    34 
       
    35 fun proofstep_to_string Axiom = "Axiom()"
       
    36 |   proofstep_to_string  (Binary ((a,b), (c,d)))=
       
    37       "Binary(("^(string_of_int a)^","^(string_of_int b)^"),("^(string_of_int c)^","^(string_of_int d)^"))"
       
    38 |   proofstep_to_string (Factor (a,b,c)) =
       
    39       "Factor("^(string_of_int a)^","^(string_of_int b)^","^(string_of_int c)^")"
       
    40 |   proofstep_to_string  (Para ((a,b), (c,d)))= 
       
    41       "Para(("^(string_of_int a)^","^(string_of_int b)^"),("^(string_of_int c)^","^(string_of_int d)^"))"
       
    42 |   proofstep_to_string  (MRR ((a,b), (c,d))) =
       
    43       "MRR(("^(string_of_int a)^","^(string_of_int b)^"),("^(string_of_int c)^","^(string_of_int d)^"))"
       
    44 (*|   proofstep_to_string (Rewrite((a,b),(c,d))) =
       
    45       "Rewrite(("^(string_of_int a)^","^(string_of_int b)^"),("^(string_of_int c)^","^(string_of_int d)^"))"*)
       
    46 
       
    47 
       
    48 fun proof_to_string (num,(step,clause_strs, thmvars)) =
       
    49  (string_of_int num)^(proofstep_to_string step)^
       
    50  "["^(clause_strs_to_string clause_strs "")^"]["^(thmvars_to_string thmvars "")^"]"
       
    51  
       
    52 
       
    53 fun proofs_to_string [] str = str
       
    54 |   proofs_to_string (x::xs) str = proofs_to_string xs (str ^ proof_to_string x)
       
    55 
       
    56 
       
    57 fun init_proofstep_to_string (num, step, clause_strs) =
       
    58  (string_of_int num)^" "^(proofstep_to_string step)^" "^
       
    59  (clause_strs_to_string clause_strs "")^" "
       
    60 
       
    61 fun init_proofsteps_to_string [] str = str
       
    62 |   init_proofsteps_to_string (x::xs) str =
       
    63       init_proofsteps_to_string xs (str ^ init_proofstep_to_string x)
       
    64   
       
    65 
       
    66 
       
    67 (*** get a string representing the Isabelle ordered axioms ***)
       
    68 
       
    69 fun origAx_to_string (num,(meta,thmvars)) =
       
    70     let val clause_strs = ReconOrderClauses.get_meta_lits_bracket meta
       
    71     in
       
    72        (string_of_int num)^"OrigAxiom()["^
       
    73        (clause_strs_to_string clause_strs "")^"]["^
       
    74        (thmvars_to_string thmvars "")^"]"
       
    75     end
       
    76 
       
    77 
       
    78 fun  origAxs_to_string [] str = str
       
    79 |   origAxs_to_string (x::xs) str = 
       
    80       origAxs_to_string xs (str ^ origAx_to_string x )
       
    81 
       
    82 
       
    83 (*** get a string representing the Isabelle ordered axioms not used in the spass proof***)
       
    84 
       
    85 fun extraAx_to_string (num, (meta,thmvars)) =
       
    86    let val clause_strs = ReconOrderClauses.get_meta_lits_bracket meta
       
    87    in
       
    88       (string_of_int num)^"ExtraAxiom()["^
       
    89       (clause_strs_to_string clause_strs "")^"]"^
       
    90       "["^(thmvars_to_string thmvars "")^"]"
       
    91    end;
       
    92 
       
    93 fun extraAxs_to_string [] str = str
       
    94 |   extraAxs_to_string (x::xs) str =
       
    95       let val newstr = extraAx_to_string x 
       
    96       in
       
    97 	  extraAxs_to_string xs (str^newstr)
       
    98       end;
       
    99 
       
   100 fun is_axiom (_,Axiom,str) = true
       
   101 |   is_axiom (_,_,_) = false
       
   102 
       
   103 fun get_step_nums [] nums = nums
       
   104 |   get_step_nums (( num:int,Axiom, str)::xs) nums = get_step_nums xs (nums@[num])
       
   105 
       
   106 exception Noassoc;
       
   107 
       
   108 fun assoc_snd a [] = raise Noassoc
       
   109   | assoc_snd a ((x, y)::t) = if a = y then x else assoc_snd a t;
       
   110 
       
   111 (* change to be something using check_order  instead of a = y --> returns true if ASSERTION not raised in checkorder, false otherwise *)
       
   112 
       
   113 (*fun get_assoc_snds [] xs assocs= assocs
       
   114 |   get_assoc_snds (x::xs) ys assocs = get_assoc_snds xs ys (assocs@[((assoc_snd x ys))])
       
   115 *)
       
   116 (*FIX - should this have vars in it? *)
       
   117 fun there_out_of_order xs ys = (ReconOrderClauses.checkorder xs ys [] ([],[],[]); true) 
       
   118                                handle _ => false
       
   119 
       
   120 fun assoc_out_of_order a [] = raise Noassoc
       
   121 |   assoc_out_of_order a ((b,c)::t) = if there_out_of_order a c then b else assoc_out_of_order a t;
       
   122 
       
   123 fun get_assoc_snds [] xs assocs= assocs
       
   124 |   get_assoc_snds (x::xs) ys assocs = get_assoc_snds xs ys (assocs@[((assoc_out_of_order x ys))])
       
   125 
       
   126 fun add_if_not_inlist eq [] xs newlist = newlist
       
   127 |   add_if_not_inlist eq (y::ys) xs newlist =
       
   128     if not (member eq xs y) then add_if_not_inlist eq ys xs (y::newlist)
       
   129     else add_if_not_inlist eq ys xs newlist
       
   130 
       
   131 (*Flattens a list of list of strings to one string*)
       
   132 fun onestr ls = String.concat (map String.concat ls);
       
   133 
       
   134 
       
   135 (**** retrieve the axioms that were obtained from the clasimpset ****)
       
   136 
       
   137 (*Get names of axioms used. Axioms are indexed from 1, while the array is indexed from 0*)
       
   138 fun get_axiom_names (names_arr: string array) step_nums = 
       
   139     let fun is_axiom n = n <= Array.length names_arr 
       
   140         fun index i = Array.sub(names_arr, i-1)
       
   141         val axnums = List.filter is_axiom step_nums
       
   142         val axnames = sort_distinct string_ord (map index axnums)
       
   143     in
       
   144 	if length axnums = length step_nums then "UNSOUND!!" :: axnames
       
   145 	else axnames
       
   146     end
       
   147 
       
   148  (*String contains multiple lines. We want those of the form 
       
   149      "253[0:Inp] et cetera..."
       
   150   A list consisting of the first number in each line is returned. *)
       
   151 fun get_spass_linenums proofstr = 
       
   152   let val toks = String.tokens (not o Char.isAlphaNum)
       
   153       fun inputno (ntok::"0"::"Inp"::_) = Int.fromString ntok
       
   154         | inputno _ = NONE
       
   155       val lines = String.tokens (fn c => c = #"\n") proofstr
       
   156   in  List.mapPartial (inputno o toks) lines  end
       
   157 
       
   158 fun get_axiom_names_spass proofstr names_arr =
       
   159    get_axiom_names names_arr (get_spass_linenums proofstr);
       
   160     
       
   161  (*String contains multiple lines. We want those of the form 
       
   162    "number : ...: ...: initial..." *)
       
   163 fun get_e_linenums proofstr = 
       
   164   let val fields = String.fields (fn c => c = #":")
       
   165       val nospaces = String.translate (fn c => if c = #" " then "" else str c)
       
   166       fun initial s = String.isPrefix "initial" (nospaces s)
       
   167       fun firstno (line :: _ :: _ :: rule :: _) = 
       
   168             if initial rule then Int.fromString line else NONE
       
   169         | firstno _ = NONE
       
   170       val lines = String.tokens (fn c => c = #"\n") proofstr
       
   171   in  List.mapPartial (firstno o fields) lines  end
       
   172 
       
   173 fun get_axiom_names_e proofstr names_arr =
       
   174    get_axiom_names names_arr (get_e_linenums proofstr);
       
   175     
       
   176  (*String contains multiple lines. We want those of the form 
       
   177      "*********** [448, input] ***********".
       
   178   A list consisting of the first number in each line is returned. *)
       
   179 fun get_vamp_linenums proofstr = 
       
   180   let val toks = String.tokens (not o Char.isAlphaNum)
       
   181       fun inputno [ntok,"input"] = Int.fromString ntok
       
   182         | inputno _ = NONE
       
   183       val lines = String.tokens (fn c => c = #"\n") proofstr
       
   184   in  List.mapPartial (inputno o toks) lines  end
       
   185 
       
   186 fun get_axiom_names_vamp proofstr names_arr =
       
   187    get_axiom_names names_arr (get_vamp_linenums proofstr);
       
   188     
       
   189 
       
   190 (***********************************************)
       
   191 (* get axioms for reconstruction               *)
       
   192 (***********************************************)
       
   193 fun numclstr (vars, []) str = str
       
   194 |   numclstr ( vars, ((num, thm)::rest)) str =
       
   195       let val newstr = str^(string_of_int num)^" "^(string_of_thm thm)^" "
       
   196       in
       
   197         numclstr  (vars,rest) newstr
       
   198       end
       
   199 
       
   200 fun addvars c (a,b)  = (a,b,c)
       
   201 
       
   202 fun get_axioms_used proof_steps thms names_arr  =
       
   203  let val axioms = (List.filter is_axiom) proof_steps
       
   204      val step_nums = get_step_nums axioms []
       
   205 
       
   206      val clauses = make_clauses thms    (*FIXME: must this be repeated??*)
       
   207      
       
   208      val vars = map thm_varnames clauses
       
   209     
       
   210      val distvars = distinct (op =) (fold append vars [])
       
   211      val clause_terms = map prop_of clauses  
       
   212      val clause_frees = List.concat (map term_frees clause_terms)
       
   213 
       
   214      val frees = map lit_string_with_nums clause_frees;
       
   215 
       
   216      val distfrees = distinct (op =) frees
       
   217 
       
   218      val metas = map Meson.make_meta_clause clauses
       
   219      val ax_strs = map #3 axioms
       
   220 
       
   221      (* literals of -all- axioms, not just those used by spass *)
       
   222      val meta_strs = map ReconOrderClauses.get_meta_lits metas
       
   223     
       
   224      val metas_and_strs = ListPair.zip (metas,meta_strs)
       
   225      val _ = trace ("\nAxioms: " ^ onestr ax_strs)
       
   226      val _ = trace ("\nMeta_strs: " ^ onestr meta_strs)
       
   227 
       
   228      (* get list of axioms as thms with their variables *)
       
   229 
       
   230      val ax_metas = get_assoc_snds ax_strs metas_and_strs []
       
   231      val ax_vars = map thm_varnames ax_metas
       
   232      val ax_with_vars = ListPair.zip (ax_metas,ax_vars)
       
   233 
       
   234      (* get list of extra axioms as thms with their variables *)
       
   235      val extra_metas = add_if_not_inlist Thm.eq_thm metas ax_metas []
       
   236      val extra_vars = map thm_varnames extra_metas
       
   237      val extra_with_vars = if not (null extra_metas)
       
   238 			   then ListPair.zip (extra_metas,extra_vars)
       
   239 			   else []
       
   240  in
       
   241     (distfrees,distvars, extra_with_vars,ax_with_vars, ListPair.zip (step_nums,ax_metas))
       
   242  end;
       
   243                                             
       
   244 
       
   245 (*********************************************************************)
       
   246 (* Pass in spass string of proof and string version of isabelle goal *)
       
   247 (* Get out reconstruction steps as a string to be sent to Isabelle   *)
       
   248 (*********************************************************************)
       
   249 
       
   250 fun rules_to_string [] = "NONE"
       
   251   | rules_to_string xs = space_implode "  " xs
       
   252 
       
   253 
       
   254 (*The signal handler in watcher.ML must be able to read the output of this.*)
       
   255 fun prover_lemma_list_aux getax proofstr probfile toParent ppid names_arr = 
       
   256  let val _ = trace
       
   257                ("\n\nGetting lemma names. proofstr is " ^ proofstr ^
       
   258                 "\nprobfile is " ^ probfile ^
       
   259                 "  num of clauses is " ^ string_of_int (Array.length names_arr))
       
   260      val axiom_names = getax proofstr names_arr
       
   261      val ax_str = rules_to_string axiom_names
       
   262     in 
       
   263 	 trace ("\nDone. Lemma list is " ^ ax_str);
       
   264          TextIO.output (toParent, "Success. Lemmas used in automatic proof: " ^
       
   265                   ax_str ^ "\n");
       
   266 	 TextIO.output (toParent, probfile ^ "\n");
       
   267 	 TextIO.flushOut toParent;
       
   268 	 Posix.Process.kill(Posix.Process.K_PROC ppid, Posix.Signal.usr2)
       
   269     end
       
   270     handle exn => (*FIXME: exn handler is too general!*)
       
   271      (trace ("\nprover_lemma_list_aux: In exception handler: " ^ 
       
   272              Toplevel.exn_message exn);
       
   273       TextIO.output (toParent, "Translation failed for the proof: " ^ 
       
   274                      String.toString proofstr ^ "\n");
       
   275       TextIO.output (toParent, probfile);
       
   276       TextIO.flushOut toParent;
       
   277       Posix.Process.kill(Posix.Process.K_PROC ppid, Posix.Signal.usr2));
       
   278 
       
   279 val e_lemma_list = prover_lemma_list_aux get_axiom_names_e;
       
   280 
       
   281 val vamp_lemma_list = prover_lemma_list_aux get_axiom_names_vamp;
       
   282 
       
   283 val spass_lemma_list = prover_lemma_list_aux get_axiom_names_spass;
       
   284 
       
   285 
       
   286 (**** Full proof reconstruction for SPASS (not really working) ****)
       
   287 
       
   288 fun spass_reconstruct proofstr probfile toParent ppid thms names_arr = 
       
   289   let val _ = trace ("\nspass_reconstruct. Proofstr is "^proofstr)
       
   290       val tokens = #1(lex proofstr)
       
   291 
       
   292   (* parse spass proof into datatype *)
       
   293   (***********************************)
       
   294       val proof_steps = parse tokens
       
   295       val _ = trace "\nParsing finished"
       
   296     
       
   297   (************************************)
       
   298   (* recreate original subgoal as thm *)
       
   299   (************************************)
       
   300       (* get axioms as correctly numbered clauses w.r.t. the Spass proof *)
       
   301       (* need to get prems_of thm, then get right one of the prems, relating to whichever*)
       
   302       (* subgoal this is, and turn it into meta_clauses *)
       
   303       (* should prob add array and table here, so that we can get axioms*)
       
   304       (* produced from the clasimpset rather than the problem *)
       
   305       val (frees,vars,extra_with_vars ,ax_with_vars,numcls) = get_axioms_used proof_steps  thms names_arr
       
   306       
       
   307       (*val numcls_string = numclstr ( vars, numcls) ""*)
       
   308       val _ = trace "\ngot axioms"
       
   309 	
       
   310   (************************************)
       
   311   (* translate proof                  *)
       
   312   (************************************)
       
   313       val _ = trace ("\nabout to translate proof, steps: "
       
   314                        ^ (init_proofsteps_to_string proof_steps ""))
       
   315       val (newthm,proof) = translate_proof numcls  proof_steps vars
       
   316       val _ = trace ("translated proof, steps: "^(init_proofsteps_to_string proof_steps ""))
       
   317   (***************************************************)
       
   318   (* transfer necessary steps as strings to Isabelle *)
       
   319   (***************************************************)
       
   320       (* turn the proof into a string *)
       
   321       val reconProofStr = proofs_to_string proof ""
       
   322       (* do the bit for the Isabelle ordered axioms at the top *)
       
   323       val ax_nums = map #1 numcls
       
   324       val ax_strs = map ReconOrderClauses.get_meta_lits_bracket (map #2 numcls)
       
   325       val numcls_strs = ListPair.zip (ax_nums,ax_strs)
       
   326       val num_cls_vars =  map (addvars vars) numcls_strs;
       
   327       val reconIsaAxStr = origAxs_to_string (ListPair.zip (ax_nums,ax_with_vars)) ""
       
   328       
       
   329       val extra_nums = if not (null extra_with_vars) then (1 upto (length extra_with_vars))
       
   330                        else []
       
   331       val reconExtraAxStr = extraAxs_to_string ( ListPair.zip (extra_nums,extra_with_vars)) ""
       
   332       val frees_str = "["^(thmvars_to_string frees "")^"]"
       
   333       val reconstr = (frees_str^reconExtraAxStr^reconIsaAxStr^reconProofStr)
       
   334       val _ = trace ("\nReconstruction:\n" ^ reconstr)
       
   335   in 
       
   336        TextIO.output (toParent, reconstr^"\n");
       
   337        TextIO.output (toParent, probfile ^ "\n");
       
   338        TextIO.flushOut toParent;
       
   339        Posix.Process.kill(Posix.Process.K_PROC ppid, Posix.Signal.usr2);
       
   340        all_tac
       
   341   end
       
   342   handle exn => (*FIXME: exn handler is too general!*)
       
   343    (trace ("\nspass_reconstruct. In exception handler: " ^ Toplevel.exn_message exn);
       
   344     TextIO.output (toParent,"Translation failed for SPASS proof:"^
       
   345          String.toString proofstr ^"\n");
       
   346     TextIO.output (toParent, probfile ^ "\n");
       
   347     TextIO.flushOut toParent;
       
   348     Posix.Process.kill(Posix.Process.K_PROC ppid, Posix.Signal.usr2); all_tac)
       
   349 
       
   350 (**********************************************************************************)
       
   351 (* At other end, want to turn back into datatype so can apply reconstruct_proof.  *)
       
   352 (* This will be done by the signal handler                                        *)
       
   353 (**********************************************************************************)
       
   354 
       
   355 (* Parse in the string version of the proof steps for reconstruction *)
       
   356 (* Isar format: cl1 [BINARY 0 cl2 0];cl1 [PARAMOD 0 cl2 0]; cl1 [DEMOD 0 cl2];cl1 [FACTOR 1 2];*)
       
   357 
       
   358 
       
   359  val term_numstep =
       
   360         (number ++ (a (Other ",")) ++ number) >> (fn (a, (_, c)) => (a, c))
       
   361 
       
   362 val extraaxiomstep = (a (Word "ExtraAxiom"))++ (a (Other "(")) ++(a (Other ")"))
       
   363             >> (fn (_) => ExtraAxiom)
       
   364 
       
   365 
       
   366 
       
   367 val origaxiomstep = (a (Word "OrigAxiom"))++ (a (Other "(")) ++(a (Other ")"))
       
   368             >> (fn (_) => OrigAxiom)
       
   369 
       
   370 
       
   371  val axiomstep = (a (Word "Axiom"))++ (a (Other "(")) ++(a (Other ")"))
       
   372             >> (fn (_) => Axiom)
       
   373      
       
   374 
       
   375 
       
   376       
       
   377  val binarystep = (a (Word "Binary")) ++ (a (Other "(")) ++ (a (Other "(")) 
       
   378                    ++ term_numstep  ++ (a (Other ")")) ++ (a (Other ","))
       
   379                    ++ (a (Other "(")) ++ term_numstep ++ (a (Other ")")) ++ (a (Other ")"))
       
   380             >> (fn (_, (_, (_, (c, (_,(_,(_, (e,(_,_))))))))) => Binary (c,e))
       
   381       
       
   382 
       
   383  val parastep = (a (Word "Para")) ++ (a (Other "(")) ++ (a (Other "(")) 
       
   384                    ++ term_numstep  ++ (a (Other ")")) ++ (a (Other ","))
       
   385                    ++ (a (Other "(")) ++ term_numstep ++ (a (Other ")")) ++ (a (Other ")"))
       
   386             >> (fn (_, (_, (_, (c, (_,(_,(_, (e,(_,_))))))))) => Para(c, e))
       
   387       
       
   388  val mrrstep = (a (Word "MRR")) ++ (a (Other "(")) ++ (a (Other "(")) 
       
   389                    ++ term_numstep  ++ (a (Other ")")) ++ (a (Other ","))
       
   390                    ++ (a (Other "(")) ++ term_numstep ++ (a (Other ")")) ++ (a (Other ")"))
       
   391             >> (fn (_, (_, (_, (c, (_,(_,(_, (e,(_,_))))))))) => MRR(c, e))
       
   392       
       
   393 
       
   394  val factorstep = (a (Word "Factor")) ++ (a (Other "("))
       
   395                     ++ number ++ (a (Other ","))
       
   396                        ++ number ++ (a (Other ","))
       
   397                        ++ number ++  (a (Other ")"))
       
   398                    
       
   399             >> (fn (_, (_, (c, (_, (e,(_,(f,_))))))) =>  Factor (c,e,f))
       
   400 
       
   401 
       
   402 (*val rewritestep = (a (Word "Rewrite"))  ++ (a (Other "(")) ++ (a (Other "(")) 
       
   403                    ++ term_numstep  ++ (a (Other ")")) ++ (a (Other ","))
       
   404                    ++ (a (Other "(")) ++ term_numstep ++ (a (Other ")")) ++ (a (Other ")"))
       
   405             >> (fn (_, (_, (_, (c, (_,(_,(_, (e,(_,_))))))))) => Rewrite (c,e))*)
       
   406 
       
   407 val obviousstep = (a (Word "Obvious")) ++ (a (Other "(")) 
       
   408                    ++ term_numstep  ++ (a (Other ")")) 
       
   409             >> (fn (_, (_, (c,_))) => Obvious (c))
       
   410 
       
   411  val methodstep = extraaxiomstep || origaxiomstep || axiomstep ||binarystep || factorstep|| parastep || mrrstep || (*rewritestep ||*) obviousstep
       
   412 
       
   413 
       
   414  val number_list_step =
       
   415         ( number ++ many ((a (Other ",") ++ number)>> #2))
       
   416         >> (fn (a,b) => (a::b))
       
   417         
       
   418  val numberlist_step = a (Other "[")  ++ a (Other "]")
       
   419                         >>(fn (_,_) => ([]:int list))
       
   420                        || a (Other "[") ++ number_list_step ++ a (Other "]")
       
   421                         >>(fn (_,(a,_)) => a)
       
   422                     
       
   423 
       
   424 
       
   425 (** change this to allow P (x U) *)
       
   426  fun arglist_step input = 
       
   427    ( word ++ many word >> (fn (a, b) => (a^" "^(space_implode " " b)))
       
   428     ||word >> (fn (a) => (a)))input
       
   429                 
       
   430 
       
   431 fun literal_step input = (word ++ a (Other "(") ++ arglist_step ++  a (Other ")")
       
   432                                           >>(fn (a, (b, (c,d))) => (a^" ("^(c)^")"))
       
   433                         || arglist_step >> (fn (a) => (a)))input
       
   434                            
       
   435 
       
   436 
       
   437 (* fun term_step input = (a (Other "~") ++ arglist_step ++ a (Other "%")>> (fn (a,(b,c)) => ("~ "^b))
       
   438                      ||  arglist_step ++ a (Other "%")>> (fn (a,b) => a ))input
       
   439 *)
       
   440 
       
   441 
       
   442  fun term_step input = (a (Other "~") ++ literal_step ++ a (Other "%")>> (fn (a,(b,c)) => ("~ "^b))
       
   443                      ||  literal_step ++ a (Other "%")>> (fn (a,b) => a ))input
       
   444 
       
   445 
       
   446          
       
   447 
       
   448  val term_list_step =
       
   449         (  term_step ++ many ( term_step))
       
   450         >> (fn (a,b) => (a::b))
       
   451         
       
   452  
       
   453 val term_lists_step = a (Other "[")  ++ a (Other "]")
       
   454                         >>(fn (_,_) => ([]:string list))
       
   455                        || a (Other "[") ++ term_list_step ++ a (Other "]")
       
   456                         >>(fn (_,(a,_)) => a)
       
   457                      
       
   458 
       
   459  val linestep = number ++ methodstep ++ term_lists_step ++ term_lists_step
       
   460                 >> (fn (a, (b, (c,d))) => (a,(b,c,d)))
       
   461     
       
   462  val lines_step = many linestep
       
   463 
       
   464  val alllines_step = (term_lists_step ++ lines_step ) ++ finished >> #1
       
   465     
       
   466  val parse_step = #1 o alllines_step
       
   467 
       
   468 
       
   469  (*
       
   470 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)"
       
   471 *)
       
   472 
       
   473 (************************************************************)
       
   474 (* Construct an Isar style proof from a list of proof steps *)
       
   475 (************************************************************)
       
   476 (* want to assume all axioms, then do haves for the other clauses*)
       
   477 (* then show for the last step *)
       
   478 
       
   479 (* replace ~ by not here *)
       
   480 val change_nots = String.translate (fn c => if c = #"~" then "\\<not>" else str c);
       
   481 
       
   482 fun clstrs_to_string xs = space_implode "; " (map change_nots xs);
       
   483 
       
   484 fun thmvars_to_quantstring [] str = str
       
   485 |   thmvars_to_quantstring (x::[]) str =str^x^". "
       
   486 |   thmvars_to_quantstring (x::xs) str = thmvars_to_quantstring xs (str^(x^" "))
       
   487 
       
   488 
       
   489 fun clause_strs_to_isar clstrs [] =
       
   490       "\"\\<lbrakk>"^(clstrs_to_string clstrs)^"\\<rbrakk> \\<Longrightarrow> False\""
       
   491 |   clause_strs_to_isar clstrs thmvars =
       
   492       "\"\\<And>"^(thmvars_to_quantstring thmvars "")^
       
   493       "\\<lbrakk>"^(clstrs_to_string clstrs)^"\\<rbrakk> \\<Longrightarrow> False\""
       
   494 
       
   495 fun frees_to_isar_str clstrs = space_implode " " (map change_nots clstrs)
       
   496 
       
   497 
       
   498 (***********************************************************************)
       
   499 (* functions for producing assumptions for the Isabelle ordered axioms *)
       
   500 (***********************************************************************)
       
   501 (*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))[][]";       
       
   502 num, rule, clausestrs, vars*)
       
   503 
       
   504 
       
   505 (* assume the extra clauses - not used in Spass proof *)
       
   506 
       
   507 fun is_extraaxiom_step ( num:int,(ExtraAxiom, str, tstr)) = true
       
   508 |   is_extraaxiom_step (num, _) = false
       
   509 
       
   510 fun get_extraaxioms xs = List.filter (is_extraaxiom_step) ( xs)
       
   511 
       
   512 fun assume_isar_extraaxiom [] str  = str
       
   513 |   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 " )
       
   514 
       
   515 
       
   516 
       
   517 fun assume_isar_extraaxioms  [] = ""
       
   518 |assume_isar_extraaxioms ((numb,(step, clstrs, thmstrs))::xs) = let val str = "assume cl"^(string_of_int numb)^"': "^(clause_strs_to_isar clstrs thmstrs)^"\n" 
       
   519                                          in
       
   520                                              assume_isar_extraaxiom xs str
       
   521                                          end
       
   522 
       
   523 (* assume the Isabelle ordered clauses *)
       
   524 
       
   525 fun is_origaxiom_step ( num:int,(OrigAxiom, str, tstr)) = true
       
   526 |   is_origaxiom_step (num, _) = false
       
   527 
       
   528 fun get_origaxioms xs = List.filter (is_origaxiom_step) ( xs)
       
   529 
       
   530 fun assume_isar_origaxiom [] str  = str
       
   531 |   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 " )
       
   532 
       
   533 
       
   534 
       
   535 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" 
       
   536                                          in
       
   537                                              assume_isar_origaxiom xs str
       
   538                                          end
       
   539 
       
   540 
       
   541 
       
   542 fun is_axiom_step ( num:int,(Axiom, str, tstr)) = true
       
   543 |   is_axiom_step (num, _) = false
       
   544 
       
   545 fun get_axioms xs = List.filter  (is_axiom_step) ( xs)
       
   546 
       
   547 fun have_isar_axiomline (numb,(step, clstrs, thmstrs))="have cl"^(string_of_int numb)^": "^(clause_strs_to_isar clstrs thmstrs)^"\n"
       
   548 
       
   549 fun  by_isar_axiomline (numb,(step, clstrs, thmstrs))="by (rule cl"^ (string_of_int numb)^"') \n"
       
   550 
       
   551 
       
   552 fun isar_axiomline (numb, (step, clstrs, thmstrs))  = (have_isar_axiomline (numb,(step,clstrs, thmstrs )))^( by_isar_axiomline(numb,(step,clstrs, thmstrs )) )
       
   553 
       
   554 
       
   555 fun isar_axiomlines [] str = str
       
   556 |   isar_axiomlines (x::xs) str = isar_axiomlines xs (str^(isar_axiomline x))
       
   557 
       
   558 
       
   559 fun have_isar_line (numb,(step, clstrs, thmstrs))="have cl"^(string_of_int numb)^": "^(clause_strs_to_isar clstrs thmstrs)^"\n"
       
   560 (*FIX: ask Larry to add and mrr attribute *)
       
   561 
       
   562 fun by_isar_line ((Binary ((a,b), (c,d)))) = 
       
   563     "by(rule cl"^
       
   564 		(string_of_int a)^" [binary "^(string_of_int b)^" cl"^
       
   565 		(string_of_int c)^" "^(string_of_int d)^"])\n"
       
   566 |by_isar_line ((MRR ((a,b), (c,d)))) = 
       
   567     "by(rule cl"^
       
   568 		(string_of_int a)^" [binary "^(string_of_int b)^" cl"^
       
   569 		(string_of_int c)^" "^(string_of_int d)^"])\n"
       
   570 |   by_isar_line ( (Para ((a,b), (c,d)))) =
       
   571     "by (rule cl"^
       
   572 		(string_of_int a)^" [paramod "^(string_of_int b)^" cl"^
       
   573 		(string_of_int c)^" "^(string_of_int d)^"])\n"
       
   574 |   by_isar_line ((Factor ((a,b,c)))) = 
       
   575     "by (rule cl"^(string_of_int a)^" [factor "^(string_of_int b)^" "^
       
   576 		(string_of_int c)^" ])\n"
       
   577 (*|   by_isar_line ( (Rewrite ((a,b),(c,d)))) =
       
   578     "by (rule cl"^(string_of_int a)^" [demod "^(string_of_int b)^" "^
       
   579 		(string_of_int c)^" "^(string_of_int d)^" ])\n"*)
       
   580 |   by_isar_line ( (Obvious ((a,b)))) =
       
   581     "by (rule cl"^(string_of_int a)^" [obvious "^(string_of_int b)^" ])\n"
       
   582 
       
   583 fun isar_line (numb, (step, clstrs, thmstrs))  = (have_isar_line (numb,(step,clstrs, thmstrs )))^( by_isar_line step)
       
   584 
       
   585 
       
   586 fun isar_lines [] str = str
       
   587 |   isar_lines (x::xs) str = isar_lines xs (str^(isar_line x))
       
   588 
       
   589 fun last_isar_line (numb,( step, clstrs,thmstrs)) = "show \"False\"\n"^(by_isar_line step)
       
   590 
       
   591 
       
   592 fun to_isar_proof (frees, xs) =
       
   593     let val extraaxioms = get_extraaxioms xs
       
   594 	val extraax_num = length extraaxioms
       
   595 	val origaxioms_and_steps = Library.drop (extraax_num, xs)  
       
   596 	
       
   597 	val origaxioms = get_origaxioms origaxioms_and_steps
       
   598 	val origax_num = length origaxioms
       
   599 	val axioms_and_steps = Library.drop (origax_num + extraax_num, xs)  
       
   600 	val axioms = get_axioms axioms_and_steps
       
   601 	
       
   602 	val steps = Library.drop (origax_num, axioms_and_steps)
       
   603 	val (firststeps, laststep) = split_last steps
       
   604 	
       
   605 	val isar_proof = 
       
   606 		("show \"[your goal]\"\n")^
       
   607 		("proof (rule ccontr,skolemize, make_clauses) \n")^
       
   608 		("fix "^(frees_to_isar_str frees)^"\n")^
       
   609 		(assume_isar_extraaxioms extraaxioms)^
       
   610 		(assume_isar_origaxioms origaxioms)^
       
   611 		(isar_axiomlines axioms "")^
       
   612 		(isar_lines firststeps "")^
       
   613 		(last_isar_line laststep)^
       
   614 		("qed")
       
   615 	val _ = trace ("\nto_isar_proof returns " ^ isar_proof)
       
   616     in
       
   617 	isar_proof
       
   618     end;
       
   619 
       
   620 (* get fix vars from axioms - all Frees *)
       
   621 (* check each clause for meta-vars and /\ over them at each step*)
       
   622 
       
   623 (*******************************************************)
       
   624 (* This assumes the thm list  "numcls" is still there  *)
       
   625 (* In reality, should probably label it with an        *)
       
   626 (* ID number identifying the subgoal.  This could      *)
       
   627 (* be passed over to the watcher, e.g.  numcls25       *)
       
   628 (*******************************************************)
       
   629 
       
   630 fun apply_res_thm str  = 
       
   631   let val tokens = #1 (lex str);
       
   632       val _ = trace ("\napply_res_thm. str is: "^str^"\n")	
       
   633       val (frees,recon_steps) = parse_step tokens 
       
   634   in 
       
   635       to_isar_proof (frees, recon_steps)
       
   636   end 
       
   637 
       
   638 end;