| author | haftmann | 
| Wed, 22 Jul 2009 18:02:10 +0200 | |
| changeset 32139 | e271a64f03ff | 
| parent 32091 | 30e2ffbba718 | 
| child 32258 | d91d394c4cab | 
| permissions | -rw-r--r-- | 
| 29268 | 1 | (* Author: L C Paulson and Claire Quigley, Cambridge University Computer Laboratory *) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 2 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 3 | (***************************************************************************) | 
| 21979 
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
 paulson parents: 
21978diff
changeset | 4 | (* Code to deal with the transfer of proofs from a prover process *) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 5 | (***************************************************************************) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 6 | signature RES_RECONSTRUCT = | 
| 24425 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 paulson parents: 
24387diff
changeset | 7 | sig | 
| 25492 
4cc7976948ac
Chained theorems are no longer mentioned in metis calls and (if used) they prevent the
 paulson parents: 
25414diff
changeset | 8 | val chained_hint: string | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 9 | |
| 24425 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 paulson parents: 
24387diff
changeset | 10 | val fix_sorts: sort Vartab.table -> term -> term | 
| 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 paulson parents: 
24387diff
changeset | 11 | val invert_const: string -> string | 
| 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 paulson parents: 
24387diff
changeset | 12 | val invert_type_const: string -> string | 
| 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 paulson parents: 
24387diff
changeset | 13 | val num_typargs: Context.theory -> string -> int | 
| 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 paulson parents: 
24387diff
changeset | 14 | val make_tvar: string -> typ | 
| 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 paulson parents: 
24387diff
changeset | 15 | val strip_prefix: string -> string -> string option | 
| 25710 
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
 paulson parents: 
25492diff
changeset | 16 | val setup: Context.theory -> Context.theory | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 17 | (* extracting lemma list*) | 
| 30874 
34927a1e0ae8
reverted to explicitly check the presence of a refutation
 immler@in.tum.de parents: 
30857diff
changeset | 18 | val find_failure: string -> string option | 
| 31840 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 19 | val lemma_list: bool -> string -> | 
| 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 20 | string * string vector * (int * int) * Proof.context * Thm.thm * int -> string | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 21 | (* structured proofs *) | 
| 31840 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 22 | val structured_proof: string -> | 
| 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 23 | string * string vector * (int * int) * Proof.context * Thm.thm * int -> string | 
| 24425 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 paulson parents: 
24387diff
changeset | 24 | end; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 25 | |
| 24425 
ca97c6f3d9cd
Returning both a "one-line" proof and a structured proof
 paulson parents: 
24387diff
changeset | 26 | structure ResReconstruct : RES_RECONSTRUCT = | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 27 | struct | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 28 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 29 | val trace_path = Path.basic "atp_trace"; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 30 | |
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 31 | fun trace s = if !Output.debugging then File.append (File.tmp_path trace_path) s | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 32 | else (); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 33 | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
31866diff
changeset | 34 | fun string_of_thm ctxt = PrintMode.setmp [] (Display.string_of_thm ctxt); | 
| 26333 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
25999diff
changeset | 35 | |
| 25710 
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
 paulson parents: 
25492diff
changeset | 36 | (*For generating structured proofs: keep every nth proof line*) | 
| 26333 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
25999diff
changeset | 37 | val (modulus, modulus_setup) = Attrib.config_int "sledgehammer_modulus" 1; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 38 | |
| 25710 
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
 paulson parents: 
25492diff
changeset | 39 | (*Indicates whether to include sort information in generated proofs*) | 
| 26333 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
25999diff
changeset | 40 | val (recon_sorts, recon_sorts_setup) = Attrib.config_bool "sledgehammer_sorts" true; | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 41 | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 42 | (*Indicated whether to generate full proofs or just lemma lists - now via setup of atps*) | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 43 | (* val (full_proofs, full_proofs_setup) = Attrib.config_bool "sledgehammer_full" false; *) | 
| 26333 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
25999diff
changeset | 44 | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 45 | val setup = modulus_setup #> recon_sorts_setup; | 
| 22044 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 46 | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 47 | (**** PARSING OF TSTP FORMAT ****) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 48 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 49 | (*Syntax trees, either termlist or formulae*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 50 | datatype stree = Int of int | Br of string * stree list; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 51 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 52 | fun atom x = Br(x,[]); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 53 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 54 | fun scons (x,y) = Br("cons", [x,y]);
 | 
| 30190 | 55 | val listof = List.foldl scons (atom "nil"); | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 56 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 57 | (*Strings enclosed in single quotes, e.g. filenames*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 58 | val quoted = $$"'" |-- Scan.repeat (~$$"'") --| $$"'" >> implode; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 59 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 60 | (*Intended for $true and $false*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 61 | fun tf s = "c_" ^ str (Char.toUpper (String.sub(s,0))) ^ String.extract(s,1,NONE); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 62 | val truefalse = $$"$" |-- Symbol.scan_id >> (atom o tf); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 63 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 64 | (*Integer constants, typically proof line numbers*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 65 | fun is_digit s = Char.isDigit (String.sub(s,0)); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 66 | val integer = Scan.many1 is_digit >> (valOf o Int.fromString o implode); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 67 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 68 | (*Generalized FO terms, which include filenames, numbers, etc.*) | 
| 25999 | 69 | fun termlist x = (term ::: Scan.repeat ($$"," |-- term)) x | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 70 | and term x = (quoted >> atom || integer>>Int || truefalse || | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 71 |               Symbol.scan_id -- Scan.optional ($$"(" |-- termlist --| $$")") [] >> Br ||
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 72 |               $$"(" |-- term --| $$")" ||
 | 
| 24547 
64c20ee76bc1
Vampire structured proofs. Better parsing; some bug fixes.
 paulson parents: 
24493diff
changeset | 73 | $$"[" |-- Scan.optional termlist [] --| $$"]" >> listof) x; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 74 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 75 | fun negate t = Br("c_Not", [t]);
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 76 | fun equate (t1,t2) = Br("c_equal", [t1,t2]);
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 77 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 78 | (*Apply equal or not-equal to a term*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 79 | fun syn_equal (t, NONE) = t | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 80 | | syn_equal (t1, SOME (NONE, t2)) = equate (t1,t2) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 81 | | syn_equal (t1, SOME (SOME _, t2)) = negate (equate (t1,t2)); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 82 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 83 | (*Literals can involve negation, = and !=.*) | 
| 24547 
64c20ee76bc1
Vampire structured proofs. Better parsing; some bug fixes.
 paulson parents: 
24493diff
changeset | 84 | fun literal x = ($$"~" |-- literal >> negate || | 
| 
64c20ee76bc1
Vampire structured proofs. Better parsing; some bug fixes.
 paulson parents: 
24493diff
changeset | 85 | (term -- Scan.option (Scan.option ($$"!") --| $$"=" -- term) >> syn_equal)) x; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 86 | |
| 25999 | 87 | val literals = literal ::: Scan.repeat ($$"|" |-- literal); | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 88 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 89 | (*Clause: a list of literals separated by the disjunction sign*) | 
| 24547 
64c20ee76bc1
Vampire structured proofs. Better parsing; some bug fixes.
 paulson parents: 
24493diff
changeset | 90 | val clause = $$"(" |-- literals --| $$")" || Scan.single literal;
 | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 91 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 92 | val annotations = $$"," |-- term -- Scan.option ($$"," |-- termlist); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 93 | |
| 25718 | 94 | (*<cnf_annotated> ::= cnf(<name>,<formula_role>,<cnf_formula><annotations>). | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 95 | The <name> could be an identifier, but we assume integers.*) | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 96 | val tstp_line = (Scan.this_string "cnf" -- $$"(") |--
 | 
| 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 97 | integer --| $$"," -- Symbol.scan_id --| $$"," -- | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 98 | clause -- Scan.option annotations --| $$ ")"; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 99 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 100 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 101 | (**** INTERPRETATION OF TSTP SYNTAX TREES ****) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 102 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 103 | exception STREE of stree; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 104 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 105 | (*If string s has the prefix s1, return the result of deleting it.*) | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 106 | fun strip_prefix s1 s = | 
| 31038 | 107 | if String.isPrefix s1 s | 
| 24182 | 108 | then SOME (ResClause.undo_ascii_of (String.extract (s, size s1, NONE))) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 109 | else NONE; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 110 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 111 | (*Invert the table of translations between Isabelle and ATPs*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 112 | val type_const_trans_table_inv = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 113 | Symtab.make (map swap (Symtab.dest ResClause.type_const_trans_table)); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 114 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 115 | fun invert_type_const c = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 116 | case Symtab.lookup type_const_trans_table_inv c of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 117 | SOME c' => c' | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 118 | | NONE => c; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 119 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 120 | fun make_tvar b = TVar(("'" ^ b, 0), HOLogic.typeS);
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 121 | fun make_var (b,T) = Var((b,0),T); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 122 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 123 | (*Type variables are given the basic sort, HOL.type. Some will later be constrained | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 124 | by information from type literals, or by type inference.*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 125 | fun type_of_stree t = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 126 | case t of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 127 | Int _ => raise STREE t | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 128 | | Br (a,ts) => | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 129 | let val Ts = map type_of_stree ts | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 130 | in | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 131 | case strip_prefix ResClause.tconst_prefix a of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 132 | SOME b => Type(invert_type_const b, Ts) | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 133 | | NONE => | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 134 | if not (null ts) then raise STREE t (*only tconsts have type arguments*) | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 135 | else | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 136 | case strip_prefix ResClause.tfree_prefix a of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 137 |                     SOME b => TFree("'" ^ b, HOLogic.typeS)
 | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 138 | | NONE => | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 139 | case strip_prefix ResClause.tvar_prefix a of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 140 | SOME b => make_tvar b | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 141 | | NONE => make_tvar a (*Variable from the ATP, say X1*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 142 | end; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 143 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 144 | (*Invert the table of translations between Isabelle and ATPs*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 145 | val const_trans_table_inv = | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 146 |       Symtab.update ("fequal", "op =")
 | 
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22728diff
changeset | 147 | (Symtab.make (map swap (Symtab.dest ResClause.const_trans_table))); | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 148 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 149 | fun invert_const c = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 150 | case Symtab.lookup const_trans_table_inv c of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 151 | SOME c' => c' | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 152 | | NONE => c; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 153 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 154 | (*The number of type arguments of a constant, zero if it's monomorphic*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 155 | fun num_typargs thy s = length (Sign.const_typargs thy (s, Sign.the_const_type thy s)); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 156 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 157 | (*Generates a constant, given its type arguments*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 158 | fun const_of thy (a,Ts) = Const(a, Sign.const_instance thy (a,Ts)); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 159 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 160 | (*First-order translation. No types are known for variables. HOLogic.typeT should allow | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 161 | them to be inferred.*) | 
| 22428 | 162 | fun term_of_stree args thy t = | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 163 | case t of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 164 | Int _ => raise STREE t | 
| 22428 | 165 |     | Br ("hBOOL",[t]) => term_of_stree [] thy t  (*ignore hBOOL*)
 | 
| 166 |     | Br ("hAPP",[t,u]) => term_of_stree (u::args) thy t
 | |
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 167 | | Br (a,ts) => | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 168 | case strip_prefix ResClause.const_prefix a of | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 169 | SOME "equal" => | 
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22728diff
changeset | 170 |               list_comb(Const ("op =", HOLogic.typeT), List.map (term_of_stree [] thy) ts)
 | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 171 | | SOME b => | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 172 | let val c = invert_const b | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 173 | val nterms = length ts - num_typargs thy c | 
| 22428 | 174 | val us = List.map (term_of_stree [] thy) (List.take(ts,nterms) @ args) | 
| 175 | (*Extra args from hAPP come AFTER any arguments given directly to the | |
| 176 | constant.*) | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 177 | val Ts = List.map type_of_stree (List.drop(ts,nterms)) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 178 | in list_comb(const_of thy (c, Ts), us) end | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 179 | | NONE => (*a variable, not a constant*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 180 | let val T = HOLogic.typeT | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 181 | val opr = (*a Free variable is typically a Skolem function*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 182 | case strip_prefix ResClause.fixed_var_prefix a of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 183 | SOME b => Free(b,T) | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 184 | | NONE => | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 185 | case strip_prefix ResClause.schematic_var_prefix a of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 186 | SOME b => make_var (b,T) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 187 | | NONE => make_var (a,T) (*Variable from the ATP, say X1*) | 
| 23519 | 188 | in list_comb (opr, List.map (term_of_stree [] thy) (ts@args)) end; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 189 | |
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 190 | (*Type class literal applied to a type. Returns triple of polarity, class, type.*) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 191 | fun constraint_of_stree pol (Br("c_Not",[t])) = constraint_of_stree (not pol) t
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 192 | | constraint_of_stree pol t = case t of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 193 | Int _ => raise STREE t | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 194 | | Br (a,ts) => | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 195 | (case (strip_prefix ResClause.class_prefix a, map type_of_stree ts) of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 196 | (SOME b, [T]) => (pol, b, T) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 197 | | _ => raise STREE t); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 198 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 199 | (** Accumulate type constraints in a clause: negative type literals **) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 200 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 201 | fun addix (key,z) = Vartab.map_default (key,[]) (cons z); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 202 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 203 | fun add_constraint ((false, cl, TFree(a,_)), vt) = addix ((a,~1),cl) vt | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 204 | | add_constraint ((false, cl, TVar(ix,_)), vt) = addix (ix,cl) vt | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 205 | | add_constraint (_, vt) = vt; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 206 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 207 | (*False literals (which E includes in its proofs) are deleted*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 208 | val nofalses = filter (not o equal HOLogic.false_const); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 209 | |
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 210 | (*Final treatment of the list of "real" literals from a clause.*) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 211 | fun finish [] = HOLogic.true_const (*No "real" literals means only type information*) | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 212 | | finish lits = | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 213 | case nofalses lits of | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 214 | [] => HOLogic.false_const (*The empty clause, since we started with real literals*) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 215 | | xs => foldr1 HOLogic.mk_disj (rev xs); | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 216 | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 217 | (*Accumulate sort constraints in vt, with "real" literals in lits.*) | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 218 | fun lits_of_strees ctxt (vt, lits) [] = (vt, finish lits) | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 219 | | lits_of_strees ctxt (vt, lits) (t::ts) = | 
| 22012 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 paulson parents: 
21999diff
changeset | 220 | lits_of_strees ctxt (add_constraint (constraint_of_stree true t, vt), lits) ts | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 221 | handle STREE _ => | 
| 22428 | 222 | lits_of_strees ctxt (vt, term_of_stree [] (ProofContext.theory_of ctxt) t :: lits) ts; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 223 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 224 | (*Update TVars/TFrees with detected sort constraints.*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 225 | fun fix_sorts vt = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 226 | let fun tysubst (Type (a, Ts)) = Type (a, map tysubst Ts) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 227 | | tysubst (TVar (xi, s)) = TVar (xi, getOpt (Vartab.lookup vt xi, s)) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 228 | | tysubst (TFree (x, s)) = TFree (x, getOpt (Vartab.lookup vt (x,~1), s)) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 229 | fun tmsubst (Const (a, T)) = Const (a, tysubst T) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 230 | | tmsubst (Free (a, T)) = Free (a, tysubst T) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 231 | | tmsubst (Var (xi, T)) = Var (xi, tysubst T) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 232 | | tmsubst (t as Bound _) = t | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 233 | | tmsubst (Abs (a, T, t)) = Abs (a, tysubst T, tmsubst t) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 234 | | tmsubst (t $ u) = tmsubst t $ tmsubst u; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 235 | in fn t => if Vartab.is_empty vt then t else tmsubst t end; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 236 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 237 | (*Interpret a list of syntax trees as a clause, given by "real" literals and sort constraints. | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 238 | vt0 holds the initial sort constraints, from the conjecture clauses.*) | 
| 23519 | 239 | fun clause_of_strees ctxt vt0 ts = | 
| 22728 | 240 | let val (vt, dt) = lits_of_strees ctxt (vt0,[]) ts in | 
| 24680 | 241 | singleton (Syntax.check_terms ctxt) (TypeInfer.constrain HOLogic.boolT (fix_sorts vt dt)) | 
| 22728 | 242 | end; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 243 | |
| 29268 | 244 | fun gen_all_vars t = fold_rev Logic.all (OldTerm.term_vars t) t; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 245 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 246 | fun ints_of_stree_aux (Int n, ns) = n::ns | 
| 30190 | 247 | | ints_of_stree_aux (Br(_,ts), ns) = List.foldl ints_of_stree_aux ns ts; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 248 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 249 | fun ints_of_stree t = ints_of_stree_aux (t, []); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 250 | |
| 25086 
729f9aad1f50
Improving the propagation of type constraints for Frees
 paulson parents: 
24958diff
changeset | 251 | fun decode_tstp vt0 (name, role, ts, annots) ctxt = | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 252 | let val deps = case annots of NONE => [] | SOME (source,_) => ints_of_stree source | 
| 25086 
729f9aad1f50
Improving the propagation of type constraints for Frees
 paulson parents: 
24958diff
changeset | 253 | val cl = clause_of_strees ctxt vt0 ts | 
| 29268 | 254 | in ((name, role, cl, deps), fold Variable.declare_term (OldTerm.term_frees cl) ctxt) end; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 255 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 256 | fun dest_tstp ((((name, role), ts), annots), chs) = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 257 | case chs of | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 258 | "."::_ => (name, role, ts, annots) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 259 |         | _ => error ("TSTP line not terminated by \".\": " ^ implode chs);
 | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 260 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 261 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 262 | (** Global sort constraints on TFrees (from tfree_tcs) are positive unit clauses. **) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 263 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 264 | fun add_tfree_constraint ((true, cl, TFree(a,_)), vt) = addix ((a,~1),cl) vt | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 265 | | add_tfree_constraint (_, vt) = vt; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 266 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 267 | fun tfree_constraints_of_clauses vt [] = vt | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 268 | | tfree_constraints_of_clauses vt ([lit]::tss) = | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 269 | (tfree_constraints_of_clauses (add_tfree_constraint (constraint_of_stree true lit, vt)) tss | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 270 | handle STREE _ => (*not a positive type constraint: ignore*) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 271 | tfree_constraints_of_clauses vt tss) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 272 | | tfree_constraints_of_clauses vt (_::tss) = tfree_constraints_of_clauses vt tss; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 273 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 274 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 275 | (**** Translation of TSTP files to Isar Proofs ****) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 276 | |
| 22012 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 paulson parents: 
21999diff
changeset | 277 | fun decode_tstp_list ctxt tuples = | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 278 | let val vt0 = tfree_constraints_of_clauses Vartab.empty (map #3 tuples) | 
| 25086 
729f9aad1f50
Improving the propagation of type constraints for Frees
 paulson parents: 
24958diff
changeset | 279 | in #1 (fold_map (decode_tstp vt0) tuples ctxt) end; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 280 | |
| 23519 | 281 | (** Finding a matching assumption. The literals may be permuted, and variable names | 
| 31038 | 282 | may disagree. We have to try all combinations of literals (quadratic!) and | 
| 23519 | 283 | match up the variable names consistently. **) | 
| 284 | ||
| 31038 | 285 | fun strip_alls_aux n (Const("all",_)$Abs(a,T,t))  =
 | 
| 23519 | 286 | strip_alls_aux (n+1) (subst_bound (Var ((a,n), T), t)) | 
| 287 | | strip_alls_aux _ t = t; | |
| 288 | ||
| 289 | val strip_alls = strip_alls_aux 0; | |
| 290 | ||
| 291 | exception MATCH_LITERAL; | |
| 22012 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 paulson parents: 
21999diff
changeset | 292 | |
| 23519 | 293 | (*Ignore types: they are not to be trusted...*) | 
| 294 | fun match_literal (t1$u1) (t2$u2) env = | |
| 295 | match_literal t1 t2 (match_literal u1 u2 env) | |
| 31038 | 296 | | match_literal (Abs (_,_,t1)) (Abs (_,_,t2)) env = | 
| 23519 | 297 | match_literal t1 t2 env | 
| 31038 | 298 | | match_literal (Bound i1) (Bound i2) env = | 
| 23519 | 299 | if i1=i2 then env else raise MATCH_LITERAL | 
| 31038 | 300 | | match_literal (Const(a1,_)) (Const(a2,_)) env = | 
| 23519 | 301 | if a1=a2 then env else raise MATCH_LITERAL | 
| 31038 | 302 | | match_literal (Free(a1,_)) (Free(a2,_)) env = | 
| 23519 | 303 | if a1=a2 then env else raise MATCH_LITERAL | 
| 304 | | match_literal (Var(ix1,_)) (Var(ix2,_)) env = insert (op =) (ix1,ix2) env | |
| 305 | | match_literal _ _ env = raise MATCH_LITERAL; | |
| 306 | ||
| 307 | (*Checking that all variable associations are unique. The list env contains no | |
| 308 | repetitions, but does it contain say (x,y) and (y,y)? *) | |
| 31038 | 309 | fun good env = | 
| 23519 | 310 | let val (xs,ys) = ListPair.unzip env | 
| 311 | in not (has_duplicates (op=) xs orelse has_duplicates (op=) ys) end; | |
| 312 | ||
| 313 | (*Match one list of literals against another, ignoring types and the order of | |
| 314 | literals. Sorting is unreliable because we don't have types or variable names.*) | |
| 315 | fun matches_aux _ [] [] = true | |
| 316 | | matches_aux env (lit::lits) ts = | |
| 317 | let fun match1 us [] = false | |
| 318 | | match1 us (t::ts) = | |
| 319 | let val env' = match_literal lit t env | |
| 31038 | 320 | in (good env' andalso matches_aux env' lits (us@ts)) orelse | 
| 321 | match1 (t::us) ts | |
| 23519 | 322 | end | 
| 323 | handle MATCH_LITERAL => match1 (t::us) ts | |
| 31038 | 324 | in match1 [] ts end; | 
| 23519 | 325 | |
| 326 | (*Is this length test useful?*) | |
| 31038 | 327 | fun matches (lits1,lits2) = | 
| 328 | length lits1 = length lits2 andalso | |
| 23519 | 329 | matches_aux [] (map Envir.eta_contract lits1) (map Envir.eta_contract lits2); | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 330 | |
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 331 | fun permuted_clause t = | 
| 24958 | 332 | let val lits = HOLogic.disjuncts t | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 333 | fun perm [] = NONE | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 334 | | perm (ctm::ctms) = | 
| 24958 | 335 | if matches (lits, HOLogic.disjuncts (HOLogic.dest_Trueprop (strip_alls ctm))) | 
| 23519 | 336 | then SOME ctm else perm ctms | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 337 | in perm end; | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 338 | |
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 339 | fun have_or_show "show " lname = "show \"" | 
| 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 340 | | have_or_show have lname = have ^ lname ^ ": \"" | 
| 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 341 | |
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 342 | (*ctms is a list of conjecture clauses as yielded by Isabelle. Those returned by the | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 343 | ATP may have their literals reordered.*) | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 344 | fun isar_lines ctxt ctms = | 
| 28572 | 345 | let val string_of = PrintMode.setmp [] (fn term => Syntax.string_of_term ctxt term) | 
| 26333 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
25999diff
changeset | 346 |       val _ = trace ("\n\nisar_lines: start\n")
 | 
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 347 | fun doline have (lname, t, []) = (*No deps: it's a conjecture clause, with no proof.*) | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 348 | (case permuted_clause t ctms of | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 349 | SOME u => "assume " ^ lname ^ ": \"" ^ string_of u ^ "\"\n" | 
| 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 350 | | NONE => "assume? " ^ lname ^ ": \"" ^ string_of t ^ "\"\n") (*no match!!*) | 
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 351 | | doline have (lname, t, deps) = | 
| 23519 | 352 | have_or_show have lname ^ string_of (gen_all_vars (HOLogic.mk_Trueprop t)) ^ | 
| 22372 | 353 | "\"\n by (metis " ^ space_implode " " deps ^ ")\n" | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 354 | fun dolines [(lname, t, deps)] = [doline "show " (lname, t, deps)] | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 355 | | dolines ((lname, t, deps)::lines) = doline "have " (lname, t, deps) :: dolines lines | 
| 25710 
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
 paulson parents: 
25492diff
changeset | 356 | in setmp show_sorts (Config.get ctxt recon_sorts) dolines end; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 357 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 358 | fun notequal t (_,t',_) = not (t aconv t'); | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 359 | |
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 360 | (*No "real" literals means only type information*) | 
| 23519 | 361 | fun eq_types t = t aconv HOLogic.true_const; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 362 | |
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22728diff
changeset | 363 | fun replace_dep (old:int, new) dep = if dep=old then new else [dep]; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 364 | |
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 365 | fun replace_deps (old:int, new) (lno, t, deps) = | 
| 30190 | 366 | (lno, t, List.foldl (op union_int) [] (map (replace_dep (old, new)) deps)); | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 367 | |
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 368 | (*Discard axioms; consolidate adjacent lines that prove the same clause, since they differ | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 369 | only in type information.*) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 370 | fun add_prfline ((lno, "axiom", t, []), lines) = (*axioms are not proof lines*) | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 371 | if eq_types t (*must be clsrel/clsarity: type information, so delete refs to it*) | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 372 | then map (replace_deps (lno, [])) lines | 
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 373 | else | 
| 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 374 | (case take_prefix (notequal t) lines of | 
| 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 375 | (_,[]) => lines (*no repetition of proof line*) | 
| 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 376 | | (pre, (lno',t',deps')::post) => (*repetition: replace later line by earlier one*) | 
| 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 377 | pre @ map (replace_deps (lno', [lno])) post) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 378 | | add_prfline ((lno, role, t, []), lines) = (*no deps: conjecture clause*) | 
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 379 | (lno, t, []) :: lines | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 380 | | add_prfline ((lno, role, t, deps), lines) = | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 381 | if eq_types t then (lno, t, deps) :: lines | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 382 | (*Type information will be deleted later; skip repetition test.*) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 383 | else (*FIXME: Doesn't this code risk conflating proofs involving different types??*) | 
| 22044 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 384 | case take_prefix (notequal t) lines of | 
| 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 385 | (_,[]) => (lno, t, deps) :: lines (*no repetition of proof line*) | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 386 | | (pre, (lno',t',deps')::post) => | 
| 22044 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 387 | (lno, t', deps) :: (*repetition: replace later line by earlier one*) | 
| 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 388 | (pre @ map (replace_deps (lno', [lno])) post); | 
| 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 389 | |
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 390 | (*Recursively delete empty lines (type information) from the proof.*) | 
| 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 391 | fun add_nonnull_prfline ((lno, t, []), lines) = (*no dependencies, so a conjecture clause*) | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 392 | if eq_types t (*must be type information, tfree_tcs, clsrel, clsarity: delete refs to it*) | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 393 | then delete_dep lno lines | 
| 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 394 | else (lno, t, []) :: lines | 
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 395 | | add_nonnull_prfline ((lno, t, deps), lines) = (lno, t, deps) :: lines | 
| 30190 | 396 | and delete_dep lno lines = List.foldr add_nonnull_prfline [] (map (replace_deps (lno, [])) lines); | 
| 22470 
0d52e5157124
No label on "show"; tries to remove dependencies more cleanly
 paulson parents: 
22428diff
changeset | 397 | |
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24920diff
changeset | 398 | fun bad_free (Free (a,_)) = String.isPrefix "sko_" a | 
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22728diff
changeset | 399 | | bad_free _ = false; | 
| 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22728diff
changeset | 400 | |
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 401 | (*TVars are forbidden in goals. Also, we don't want lines with <2 dependencies. | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 402 | To further compress proofs, setting modulus:=n deletes every nth line, and nlines | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 403 | counts the number of proof lines processed so far. | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 404 | Deleted lines are replaced by their own dependencies. Note that the "add_nonnull_prfline" | 
| 22044 
6c0702a96076
More compact proof reconstruction: lines having fewer than !min_deps dependences are folded
 paulson parents: 
22012diff
changeset | 405 | phase may delete some dependencies, hence this phase comes later.*) | 
| 25710 
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
 paulson parents: 
25492diff
changeset | 406 | fun add_wanted_prfline ctxt ((lno, t, []), (nlines, lines)) = | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 407 | (nlines, (lno, t, []) :: lines) (*conjecture clauses must be kept*) | 
| 25710 
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
 paulson parents: 
25492diff
changeset | 408 | | add_wanted_prfline ctxt ((lno, t, deps), (nlines, lines)) = | 
| 29272 | 409 | if eq_types t orelse not (null (Term.add_tvars t [])) orelse | 
| 29268 | 410 | exists_subterm bad_free t orelse | 
| 24937 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 paulson parents: 
24920diff
changeset | 411 | (not (null lines) andalso (*final line can't be deleted for these reasons*) | 
| 31038 | 412 | (length deps < 2 orelse nlines mod (Config.get ctxt modulus) <> 0)) | 
| 22491 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 413 | then (nlines+1, map (replace_deps (lno, deps)) lines) (*Delete line*) | 
| 
535fbed859da
Numerous bug fixes. Type clauses distinguished from empty clauses. Working proof reduction.
 paulson parents: 
22470diff
changeset | 414 | else (nlines+1, (lno, t, deps) :: lines); | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 415 | |
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 416 | (*Replace numeric proof lines by strings, either from thm_names or sequential line numbers*) | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 417 | fun stringify_deps thm_names deps_map [] = [] | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 418 | | stringify_deps thm_names deps_map ((lno, t, deps) :: lines) = | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 419 | if lno <= Vector.length thm_names (*axiom*) | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 420 | then (Vector.sub(thm_names,lno-1), t, []) :: stringify_deps thm_names deps_map lines | 
| 21979 
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
 paulson parents: 
21978diff
changeset | 421 | else let val lname = Int.toString (length deps_map) | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 422 | fun fix lno = if lno <= Vector.length thm_names | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 423 | then SOME(Vector.sub(thm_names,lno-1)) | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 424 | else AList.lookup op= deps_map lno; | 
| 22731 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 paulson parents: 
22728diff
changeset | 425 | in (lname, t, List.mapPartial fix (distinct (op=) deps)) :: | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 426 | stringify_deps thm_names ((lno,lname)::deps_map) lines | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 427 | end; | 
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 428 | |
| 24547 
64c20ee76bc1
Vampire structured proofs. Better parsing; some bug fixes.
 paulson parents: 
24493diff
changeset | 429 | val proofstart = "proof (neg_clausify)\n"; | 
| 21979 
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
 paulson parents: 
21978diff
changeset | 430 | |
| 
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
 paulson parents: 
21978diff
changeset | 431 | fun isar_header [] = proofstart | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 432 | | isar_header ts = proofstart ^ "fix " ^ space_implode " " ts ^ "\n"; | 
| 21979 
80e10f181c46
Improvements to proof reconstruction. Now "fixes" is inserted
 paulson parents: 
21978diff
changeset | 433 | |
| 22012 
adf68479ae1b
Proof.context now sent to watcher and used in type inference step of proof reconstruction
 paulson parents: 
21999diff
changeset | 434 | fun decode_tstp_file cnfs ctxt th sgno thm_names = | 
| 26333 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
25999diff
changeset | 435 | let val _ = trace "\ndecode_tstp_file: start\n" | 
| 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
25999diff
changeset | 436 | val tuples = map (dest_tstp o tstp_line o explode) cnfs | 
| 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
25999diff
changeset | 437 | val _ = trace (Int.toString (length tuples) ^ " tuples extracted\n") | 
| 24552 | 438 | val ctxt = ProofContext.set_mode ProofContext.mode_schematic ctxt | 
| 30190 | 439 | val raw_lines = List.foldr add_prfline [] (decode_tstp_list ctxt tuples) | 
| 26333 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
25999diff
changeset | 440 | val _ = trace (Int.toString (length raw_lines) ^ " raw_lines extracted\n") | 
| 30190 | 441 | val nonnull_lines = List.foldr add_nonnull_prfline [] raw_lines | 
| 26333 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
25999diff
changeset | 442 | val _ = trace (Int.toString (length nonnull_lines) ^ " nonnull_lines extracted\n") | 
| 30190 | 443 | val (_,lines) = List.foldr (add_wanted_prfline ctxt) (0,[]) nonnull_lines | 
| 26333 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
25999diff
changeset | 444 | val _ = trace (Int.toString (length lines) ^ " lines extracted\n") | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 445 | val (ccls,fixes) = ResAxioms.neg_conjecture_clauses th sgno | 
| 26333 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
25999diff
changeset | 446 | val _ = trace (Int.toString (length ccls) ^ " conjecture clauses\n") | 
| 21999 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 paulson parents: 
21979diff
changeset | 447 | val ccls = map forall_intr_vars ccls | 
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
31866diff
changeset | 448 | val _ = | 
| 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
31866diff
changeset | 449 |         if !Output.debugging then app (fn th => trace ("\nccl: " ^ string_of_thm ctxt th)) ccls
 | 
| 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
31866diff
changeset | 450 | else () | 
| 26333 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
25999diff
changeset | 451 | val ilines = isar_lines ctxt (map prop_of ccls) (stringify_deps thm_names [] lines) | 
| 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
25999diff
changeset | 452 | val _ = trace "\ndecode_tstp_file: finishing\n" | 
| 23139 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 wenzelm parents: 
23083diff
changeset | 453 | in | 
| 26333 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
25999diff
changeset | 454 | isar_header (map #1 fixes) ^ String.concat ilines ^ "qed\n" | 
| 24547 
64c20ee76bc1
Vampire structured proofs. Better parsing; some bug fixes.
 paulson parents: 
24493diff
changeset | 455 | end | 
| 
64c20ee76bc1
Vampire structured proofs. Better parsing; some bug fixes.
 paulson parents: 
24493diff
changeset | 456 | handle e => (*FIXME: exn handler is too general!*) | 
| 31479 | 457 | let val msg = "Translation of TSTP raised an exception: " ^ ML_Compiler.exn_message e | 
| 26333 
68e5eee47a45
Attributes sledgehammer_full, sledgehammer_modulus, sledgehammer_sorts
 paulson parents: 
25999diff
changeset | 458 | in trace msg; msg end; | 
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 459 | |
| 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 460 | |
| 31866 | 461 | (*=== EXTRACTING PROOF-TEXT === *) | 
| 462 | ||
| 463 | val begin_proof_strings = ["# SZS output start CNFRefutation.", | |
| 464 | "=========== Refutation ==========", | |
| 465 | "Here is a proof"]; | |
| 466 | val end_proof_strings = ["# SZS output end CNFRefutation", | |
| 467 | "======= End of refutation =======", | |
| 468 | "Formulae used in the proof"]; | |
| 469 | fun get_proof_extract proof = | |
| 470 | let | |
| 471 | (*splits to_split by the first possible of a list of splitters*) | |
| 472 | val (begin_string, end_string) = | |
| 473 | (find_first (fn s => String.isSubstring s proof) begin_proof_strings, | |
| 474 | find_first (fn s => String.isSubstring s proof) end_proof_strings) | |
| 475 | in | |
| 476 | if is_none begin_string orelse is_none end_string | |
| 477 | then error "Could not extract proof (no substring indicating a proof)" | |
| 478 | else proof |> first_field (the begin_string) |> the |> snd | |
| 479 | |> first_field (the end_string) |> the |> fst end; | |
| 480 | ||
| 481 | (* ==== CHECK IF PROOF OF E OR VAMPIRE WAS SUCCESSFUL === *) | |
| 482 | ||
| 30874 
34927a1e0ae8
reverted to explicitly check the presence of a refutation
 immler@in.tum.de parents: 
30857diff
changeset | 483 | val failure_strings_E = ["SZS status: Satisfiable","SZS status Satisfiable", | 
| 
34927a1e0ae8
reverted to explicitly check the presence of a refutation
 immler@in.tum.de parents: 
30857diff
changeset | 484 | "SZS status: ResourceOut","SZS status ResourceOut","# Cannot determine problem status"]; | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 485 | val failure_strings_vampire = ["Satisfiability detected", "Refutation not found", "CANNOT PROVE"]; | 
| 30874 
34927a1e0ae8
reverted to explicitly check the presence of a refutation
 immler@in.tum.de parents: 
30857diff
changeset | 486 | val failure_strings_SPASS = ["SPASS beiseite: Completion found.", | 
| 
34927a1e0ae8
reverted to explicitly check the presence of a refutation
 immler@in.tum.de parents: 
30857diff
changeset | 487 | "SPASS beiseite: Ran out of time.", "SPASS beiseite: Maximal number of loops exceeded."]; | 
| 
34927a1e0ae8
reverted to explicitly check the presence of a refutation
 immler@in.tum.de parents: 
30857diff
changeset | 488 | val failure_strings_remote = ["Remote-script could not extract proof"]; | 
| 
34927a1e0ae8
reverted to explicitly check the presence of a refutation
 immler@in.tum.de parents: 
30857diff
changeset | 489 | fun find_failure proof = | 
| 29597 | 490 | let val failures = | 
| 31038 | 491 | map_filter (fn s => if String.isSubstring s proof then SOME s else NONE) | 
| 31866 | 492 | (failure_strings_E @ failure_strings_vampire @ failure_strings_SPASS @ failure_strings_remote) | 
| 493 | val correct = null failures andalso | |
| 494 | exists (fn s => String.isSubstring s proof) begin_proof_strings andalso | |
| 495 | exists (fn s => String.isSubstring s proof) end_proof_strings | |
| 496 | in | |
| 497 | if correct then NONE | |
| 498 | else if null failures then SOME "Output of ATP not in proper format" | |
| 499 | else SOME (hd failures) end; | |
| 500 | ||
| 31038 | 501 | (* === EXTRACTING LEMMAS === *) | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 502 | (* lines have the form "cnf(108, axiom, ...", | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 503 | the number (108) has to be extracted)*) | 
| 31840 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 504 | fun get_step_nums false proofextract = | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 505 | let val toks = String.tokens (not o Char.isAlphaNum) | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 506 |     fun inputno ("cnf"::ntok::"axiom"::_) = Int.fromString ntok
 | 
| 31840 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 507 |       | inputno ("cnf"::ntok::"negated"::"conjecture"::_) = Int.fromString ntok
 | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 508 | | inputno _ = NONE | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 509 | val lines = split_lines proofextract | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 510 | in List.mapPartial (inputno o toks) lines end | 
| 31840 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 511 | (*String contains multiple lines. We want those of the form | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 512 | "253[0:Inp] et cetera..." | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 513 | A list consisting of the first number in each line is returned. *) | 
| 31840 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 514 | | get_step_nums true proofextract = | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 515 | let val toks = String.tokens (not o Char.isAlphaNum) | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 516 | fun inputno (ntok::"0"::"Inp"::_) = Int.fromString ntok | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 517 | | inputno _ = NONE | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 518 | val lines = split_lines proofextract | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 519 | in List.mapPartial (inputno o toks) lines end | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 520 | |
| 31038 | 521 | (*extracting lemmas from tstp-output between the lines from above*) | 
| 31840 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 522 | fun extract_lemmas get_step_nums (proof, thm_names, conj_count, _, _, _) = | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 523 | let | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 524 | (* get the names of axioms from their numbers*) | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 525 | fun get_axiom_names thm_names step_nums = | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 526 | let | 
| 31840 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 527 | val last_axiom = Vector.length thm_names | 
| 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 528 | fun is_axiom n = n <= last_axiom | 
| 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 529 | fun is_conj n = n >= #1 conj_count andalso n < #1 conj_count + #2 conj_count | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 530 | fun getname i = Vector.sub(thm_names, i-1) | 
| 31038 | 531 | in | 
| 31840 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 532 | (sort_distinct string_ord (filter (fn x => x <> "??.unknown") | 
| 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 533 | (map getname (filter is_axiom step_nums))), | 
| 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 534 | exists is_conj step_nums) | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 535 | end | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 536 | val proofextract = get_proof_extract proof | 
| 31038 | 537 | in | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 538 | get_axiom_names thm_names (get_step_nums proofextract) | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 539 | end; | 
| 31410 | 540 | |
| 541 | (*Used to label theorems chained into the sledgehammer call*) | |
| 542 | val chained_hint = "CHAINED"; | |
| 543 | val nochained = filter_out (fn y => y = chained_hint) | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
27865diff
changeset | 544 | |
| 31038 | 545 | (* metis-command *) | 
| 546 | fun metis_line [] = "apply metis" | |
| 547 | | metis_line xs = "apply (metis " ^ space_implode " " xs ^ ")" | |
| 548 | ||
| 549 | (* atp_minimize [atp=<prover>] <lemmas> *) | |
| 550 | fun minimize_line _ [] = "" | |
| 551 | | minimize_line name lemmas = "For minimizing the number of lemmas try this command:\n" ^ | |
| 31410 | 552 |           (Markup.markup Markup.sendback) ("atp_minimize [atp=" ^ name ^ "] " ^
 | 
| 553 | space_implode " " (nochained lemmas)) | |
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: 
30874diff
changeset | 554 | |
| 31038 | 555 | fun sendback_metis_nochained lemmas = | 
| 31410 | 556 | (Markup.markup Markup.sendback o metis_line) (nochained lemmas) | 
| 31840 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 557 | |
| 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 558 | fun lemma_list dfg name result = | 
| 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 559 | let val (lemmas, used_conj) = extract_lemmas (get_step_nums dfg) result | 
| 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 560 | in sendback_metis_nochained lemmas ^ "\n" ^ minimize_line name lemmas ^ | 
| 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 561 | (if used_conj then "" | 
| 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 562 | else "\nWarning: Goal is provable because context is inconsistent.") | 
| 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 563 | end; | 
| 31037 
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
 immler@in.tum.de parents: 
30874diff
changeset | 564 | |
| 31038 | 565 | (* === Extracting structured Isar-proof === *) | 
| 31840 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 566 | fun structured_proof name (result as (proof, thm_names, conj_count, ctxt, goal, subgoalno)) = | 
| 31038 | 567 | let | 
| 568 | (*Could use split_lines, but it can return blank lines...*) | |
| 569 | val lines = String.tokens (equal #"\n"); | |
| 570 | val nospaces = String.translate (fn c => if Char.isSpace c then "" else str c) | |
| 571 | val proofextract = get_proof_extract proof | |
| 572 |     val cnfs = filter (String.isPrefix "cnf(") (map nospaces (lines proofextract))
 | |
| 31840 
beeaa1ed1f47
check if conjectures have been used in proof
 immler@in.tum.de parents: 
31479diff
changeset | 573 | val one_line_proof = lemma_list false name result | 
| 31038 | 574 | val structured = if chained_hint mem_string (String.tokens (fn c => c = #" ") one_line_proof) then "" | 
| 575 | else decode_tstp_file cnfs ctxt goal subgoalno thm_names | |
| 576 | in | |
| 577 | one_line_proof ^ "\n\n" ^ (Markup.markup Markup.sendback) structured | |
| 578 | end | |
| 21978 
72c21698a055
Contains old Tools/ATP/AtpCommunication.ML, plus proof reconstruction
 paulson parents: diff
changeset | 579 | |
| 31038 | 580 | end; |