| author | blanchet | 
| Wed, 03 Jan 2018 11:06:13 +0100 | |
| changeset 67332 | cb96edae56ef | 
| parent 64560 | c48becd96398 | 
| permissions | -rw-r--r-- | 
| 47412 | 1 | (* Title: HOL/TPTP/TPTP_Parser/tptp_to_dot.ML | 
| 2 | Author: Nik Sultana, Cambridge University Computer Laboratory | |
| 3 | ||
| 4 | Translates parsed TPTP proofs into DOT format. This can then be processed | |
| 5 | by an accompanying script to translate the proofs into other formats. | |
| 55586 | 6 | |
| 7 | It tries to adhere to the symbols used in IDV, as described in | |
| 8 | "An Interactive Derivation Viewer" by Trac & Puzis & Sutcliffe, UITP 2006. | |
| 47412 | 9 | *) | 
| 10 | ||
| 11 | signature TPTP_TO_DOT = | |
| 12 | sig | |
| 13 | (*DOT-drawing function, works directly on parsed TPTP*) | |
| 14 | val tptp_dot_node : bool -> bool -> TPTP_Syntax.tptp_line -> string | |
| 15 | ||
| 16 | (*Parse a (LEO-II+E) proof and produce a DOT file*) | |
| 17 | val write_proof_dot : string -> string -> unit | |
| 18 | end | |
| 19 | ||
| 20 | structure TPTP_To_Dot : TPTP_TO_DOT = | |
| 21 | struct | |
| 22 | ||
| 23 | open TPTP_Syntax | |
| 24 | ||
| 55592 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 25 | datatype style = | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 26 | (*Only draw shapes. No formulas or edge labels.*) | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 27 | Shapes | 
| 55593 | 28 | (*Don't draw shapes. Only write formulas (as nodes) and inference names (as edge labels).*) | 
| 55592 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 29 | | Formulas | 
| 55593 | 30 | (*Draw shapes and write the AF ID inside.*) | 
| 31 | | IDs | |
| 55592 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 32 | |
| 55594 | 33 | (*FIXME this kind of configurability isn't very user-friendly. | 
| 34 | Ideally we'd accept a parameter from the tptp_graph script.*) | |
| 55592 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 35 | (*Determine the require output style form the TPTP_GRAPH environment variable. | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 36 | Shapes is the default style.*) | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 37 | val required_style = | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 38 | if getenv "TPTP_GRAPH" = "formulas" then | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 39 | Formulas | 
| 55593 | 40 | else if getenv "TPTP_GRAPH" = "IDs" then | 
| 41 | IDs | |
| 55592 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 42 | else Shapes | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 43 | |
| 47412 | 44 | (*Draw an arc between two nodes*) | 
| 45 | fun dot_arc reverse (src, label) target = | |
| 55592 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 46 | let | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 47 | val edge_label = | 
| 55593 | 48 | if required_style = Shapes orelse required_style = IDs then "" | 
| 55592 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 49 | else | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 50 | case label of | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 51 | NONE => "" | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 52 | | SOME label => "[label=\"" ^ label ^ "\"];" | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 53 | in | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 54 | "\"" ^ (if reverse then target else src) ^ | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 55 | "\" -> \"" ^ (if reverse then src else target) ^ | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 56 | "\" " ^ edge_label ^ "\n" | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 57 | end | 
| 47412 | 58 | |
| 59 | (*Node shapes indicate the role of the related clauses.*) | |
| 60 | exception NO_ROLE_SHAPE | |
| 61 | fun the_role_shape role = | |
| 55592 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 62 | if role = Role_Fi_Domain orelse | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 63 | role = Role_Fi_Functors orelse | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 64 | role = Role_Fi_Predicates orelse | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 65 | role = Role_Type orelse | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 66 | role = Role_Unknown then | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 67 | raise NO_ROLE_SHAPE | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 68 | else if required_style = Formulas then "plaintext" | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 69 | else | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 70 | case role of | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 71 | Role_Axiom => "triangle" | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 72 | | Role_Hypothesis => "invtrapezium" | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 73 | | Role_Definition => "invtriangle" (*NOTE this is not standard wrt IDV*) | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 74 | | Role_Assumption => "trapezium" (*NOTE this is not standard wrt IDV*) | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 75 | | Role_Lemma => "hexagon" | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 76 | | Role_Theorem => "star" (*NOTE this is not standard wrt IDV*) | 
| 55590 | 77 | |
| 55592 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 78 | | Role_Conjecture => "house" | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 79 | | Role_Negated_Conjecture => "invhouse" | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 80 | | Role_Plain => "circle" | 
| 55590 | 81 | |
| 47412 | 82 | fun have_role_shape role = | 
| 83 | (the_role_shape role; true) | |
| 84 | handle NO_ROLE_SHAPE => false | |
| 85 | | exc => raise exc | |
| 86 | ||
| 87 | (*Different styles are applied to nodes relating to clauses written in | |
| 88 | difference languages.*) | |
| 89 | fun the_lang_style lang = | |
| 90 | case lang of | |
| 91 | CNF => "dotted" | |
| 92 | | FOF => "dashed" | |
| 57280 
6907432c47b9
made 'tptp_graph' more liberal (why reject TFF?)
 blanchet parents: 
55594diff
changeset | 93 | | _ => "" | 
| 47412 | 94 | |
| 55594 | 95 | (*Check if the formula just consists of "$false"? | 
| 96 | which we interpret to be the last line of the proof.*) | |
| 47412 | 97 | fun is_last_line CNF (Pred (Interpreted_Logic False, [])) = true | 
| 98 | | is_last_line THF (Atom (THF_Atom_term | |
| 64560 
c48becd96398
support THF1 parsing for 'tptp_...' commands (e.g. 'tptp_isabelle') -- useful for CASC, SystemOnTPTP, ...
 blanchet parents: 
57280diff
changeset | 99 | (Term_Func (Interpreted_Logic False, [], [])))) = true | 
| 47412 | 100 | | is_last_line _ _ = false | 
| 101 | ||
| 102 | fun tptp_dot_node with_label reverse_arrows | |
| 55592 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 103 | (Annotated_Formula (_, lang, n, role, fmla_tptp, annot)) = | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 104 | let | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 105 | val node_label = | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 106 | if required_style = Formulas then | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 107 | "\", label=\"$" ^ TPTP_Syntax.latex_of_tptp_formula fmla_tptp ^ "$\"];\n" | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 108 | (*FIXME add a parameter to switch to using the following code, which lowers, centers, and horizontally-bounds the label. | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 109 | (this is useful if you want to keep the shapes but also show formulas)*) | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 110 |    (*    "\", label=\"\\\\begin{minipage}{10cm}\\\\vspace{21mm}\\\\centering$" ^ TPTP_Syntax.latex_of_tptp_formula fmla_tptp ^ "$\\\\end{minipage}\"];\n") ^*)
 | 
| 55593 | 111 | else if required_style = IDs then | 
| 112 | "\", label=\"" ^ n ^ "\"];\n" | |
| 55592 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 113 | else | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 114 | "\", label=\"\"];\n" | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 115 | in | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 116 | (*don't expect to find 'Include' in proofs*) | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 117 | if have_role_shape role andalso role <> Role_Definition then | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 118 | "\"" ^ n ^ | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 119 | "\" [shape=\"" ^ | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 120 | (if is_last_line lang fmla_tptp then "doublecircle" | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 121 | else the_role_shape role) ^ | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 122 | "\", style=\"" ^ the_lang_style lang ^ | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 123 | node_label ^ | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 124 | (case TPTP_Proof.extract_source_info annot of | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 125 | SOME (TPTP_Proof.Inference (rule, _, pinfos)) => | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 126 | let | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 127 | fun parent_id (TPTP_Proof.Parent n) = n | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 128 | | parent_id (TPTP_Proof.ParentWithDetails (n, _)) = n | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 129 | val parent_ids = map parent_id pinfos | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 130 | in | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 131 | map | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 132 | (dot_arc reverse_arrows | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 133 | (n, if with_label then SOME rule else NONE)) | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 134 | parent_ids | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 135 | |> implode | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 136 | end | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 137 | | _ => "") | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 138 | else "" | 
| 
37c1abaf4876
behaviour of tptp_graph now depends on an envir variable to indicate whether to produce proof graph containing formulas or not;
 sultana parents: 
55591diff
changeset | 139 | end | 
| 47412 | 140 | |
| 141 | (*FIXME add opts to label arcs etc*) | |
| 142 | fun write_proof_dot input_file output_file = | |
| 143 | let | |
| 55594 | 144 | (*NOTE sometimes useful to include: rankdir=\"LR\";\n*) | 
| 47412 | 145 | val defaults = | 
| 55587 
5d3db2c626e3
experimenting with improving DOT output, and embedding LaTeX code for formulas (rather than only giving the clause number);
 sultana parents: 
55586diff
changeset | 146 | "graph[nodesep=3];\n" ^ | 
| 47412 | 147 | "node[fixedsize=true];\n" ^ | 
| 55587 
5d3db2c626e3
experimenting with improving DOT output, and embedding LaTeX code for formulas (rather than only giving the clause number);
 sultana parents: 
55586diff
changeset | 148 | "node[width=0.5];\n" ^ | 
| 47412 | 149 | "node[shape=plaintext];\n" ^ | 
| 55594 | 150 | (*NOTE sometimes useful to include: "node[fillcolor=lightgray];\n" ^ *) | 
| 55590 | 151 | "node[fontsize=50];\n" | 
| 47412 | 152 | in | 
| 153 | TPTP_Parser.parse_file input_file | |
| 55587 
5d3db2c626e3
experimenting with improving DOT output, and embedding LaTeX code for formulas (rather than only giving the clause number);
 sultana parents: 
55586diff
changeset | 154 | |> map (tptp_dot_node true true) | 
| 47426 | 155 | |> implode | 
| 47412 | 156 |     |> (fn str => "digraph ProofGraph {\n" ^ defaults ^ str ^ "}")
 | 
| 157 | |> File.write (Path.explode output_file) | |
| 158 | end | |
| 159 | ||
| 160 | end |