| author | berghofe | 
| Mon, 12 Mar 2007 11:07:59 +0100 | |
| changeset 22436 | c9e384a956df | 
| parent 22405 | 7eef90be0db4 | 
| child 23588 | 4fc6df2c7098 | 
| permissions | -rw-r--r-- | 
| 21637 | 1 | (* Title: Pure/ProofGeneral/parsing.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: David Aspinall | |
| 4 | ||
| 21867 | 5 | Parsing Isabelle theory files to add PGIP markup. | 
| 21637 | 6 | *) | 
| 7 | ||
| 8 | signature PGIP_PARSER = | |
| 9 | sig | |
| 21867 | 10 | val pgip_parser: PgipMarkup.pgip_parser | 
| 21971 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 11 | val init : unit -> unit (* clear state *) | 
| 21637 | 12 | end | 
| 13 | ||
| 14 | structure PgipParser : PGIP_PARSER = | |
| 15 | struct | |
| 16 | ||
| 17 | ||
| 18 | (** Parsing proof scripts without execution **) | |
| 19 | ||
| 20 | structure P = OuterParse; | |
| 21 | ||
| 21867 | 22 | structure D = PgipMarkup; | 
| 22405 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 23 | structure I = PgipIsabelle; | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 24 | structure T = PgipTypes; | 
| 21867 | 25 | |
| 21637 | 26 | (* Notes. | 
| 27 | This is quite tricky, because 1) we need to put back whitespace which | |
| 28 | was removed during parsing so we can provide markup around commands; | |
| 29 | 2) parsing is intertwined with execution in Isar so we have to repeat | |
| 30 | the parsing to extract interesting parts of commands. The trace of | |
| 31 | tokens parsed which is now recorded in each transition (added by | |
| 32 | Markus June '04) helps do this, but the mechanism is still a bad mess. | |
| 33 | ||
| 34 | If we had proper parse trees it would be easy, although having a | |
| 21867 | 35 | fixed type of trees might be hard with Isar's extensible parser. | 
| 21637 | 36 | |
| 37 | Probably the best solution is to produce the meta-information at | |
| 38 | the same time as the parse, for each command, e.g. by recording a | |
| 39 | list of (name,objtype) pairs which record the bindings created by | |
| 40 | a command. This would require changing the interfaces in | |
| 41 | outer_syntax.ML (or providing new ones), | |
| 42 | ||
| 43 | datatype metainfo = Binding of string * string (* name, objtype *) | |
| 44 | ||
| 45 | val command_withmetainfo: string -> string -> string -> | |
| 46 | (token list -> | |
| 47 | ((Toplevel.transition -> Toplevel.transition) * metainfo list) * | |
| 48 | token list) -> parser | |
| 49 | ||
| 50 | We'd also like to render terms as they appear in output, but this | |
| 51 | will be difficult because inner syntax is extensible and we don't | |
| 52 | have the correct syntax to hand when just doing outer parsing | |
| 53 | without execution. A reasonable approximation could | |
| 54 | perhaps be obtained by using the syntax of the current context. | |
| 55 | However, this would mean more mess trying to pick out terms, | |
| 56 | so at this stage a more serious change to the Isar functions | |
| 57 | seems necessary. | |
| 58 | *) | |
| 59 | ||
| 60 | local | |
| 21971 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 61 | val keywords_classification_table = ref (NONE: string Symtab.table option) | 
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 62 | |
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 63 | in | 
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 64 | fun get_keywords_classification_table () = | 
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 65 | case (!keywords_classification_table) of | 
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 66 | SOME t => t | 
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 67 | | NONE => (let | 
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 68 | val tab = fold (fn (c, _, k, _) => Symtab.update (c, k)) | 
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 69 | (OuterSyntax.dest_parsers ()) Symtab.empty; | 
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 70 | in (keywords_classification_table := SOME tab; tab) end) | 
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 71 | |
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 72 | (* To allow dynamic extensions to language during interactive use | 
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 73 | we need a hook in outer_syntax.ML to reset our table. But this is | 
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 74 | pretty obscure; initialisation at startup should suffice. *) | 
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 75 | |
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 76 | fun init () = (keywords_classification_table := NONE) | 
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 77 | end | 
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 78 | |
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 79 | |
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 80 | local | 
| 21637 | 81 | |
| 21867 | 82 |     fun whitespace str = D.Whitespace { text=str }
 | 
| 21637 | 83 | |
| 84 | (* an extra token is needed to terminate the command *) | |
| 85 | val sync = OuterSyntax.scan "\\<^sync>"; | |
| 86 | ||
| 21867 | 87 | fun nameparse elt objtyp nameP toks = | 
| 88 | (fst (nameP (toks@sync))) | |
| 89 |         handle _ => (error ("Error occurred in name parser for " ^ elt ^
 | |
| 22405 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 90 | "(objtype: " ^ (T.name_of_objtype objtyp) ^ ")"); | 
| 21867 | 91 | "") | 
| 21637 | 92 | |
| 93 |     fun parse_warning msg = warning ("script->PGIP markup parser: " ^ msg)
 | |
| 94 | ||
| 95 | fun xmls_of_thy_begin (name,toks,str) = (* see isar_syn.ML/theoryP *) | |
| 96 | let | |
| 97 | val ((thyname,imports,files),_) = ThyHeader.args (toks@sync) | |
| 98 | in | |
| 21867 | 99 |             [D.Opentheory { thyname=thyname, 
 | 
| 100 | parentnames = imports, | |
| 22405 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 101 | text = str }, | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 102 | 	     D.Openblock { metavarid=NONE,name=NONE,objtype=SOME I.ObjTheoryBody }]
 | 
| 21637 | 103 | end | 
| 104 | ||
| 105 | fun xmls_of_thy_decl (name,toks,str) = (* see isar_syn.ML/thy_decl cases *) | |
| 106 | let | |
| 107 | (* NB: PGIP only deals with single name bindings at the moment; | |
| 108 | potentially we could markup grouped definitions several times but | |
| 109 | that might suggest the wrong structure to the editor? | |
| 110 | Better alternative would be to put naming closer around text, | |
| 111 | but to do that we'd be much better off modifying the real parser | |
| 112 | instead of reconstructing it here. *) | |
| 113 | ||
| 114 | val plain_items = (* no names, unimportant names, or too difficult *) | |
| 115 | ["defaultsort","arities","judgement","finalconsts", | |
| 116 | "syntax", "nonterminals", "translations", | |
| 117 | "global", "local", "hide", | |
| 118 | "ML_setup", "setup", "method_setup", | |
| 119 | "parse_ast_translation", "parse_translation", "print_translation", | |
| 120 | "typed_print_translation", "print_ast_translation", "token_translation", | |
| 121 | "oracle", | |
| 122 | "declare"] | |
| 123 | ||
| 21867 | 124 | fun named_item nameP ty = | 
| 22405 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 125 | 		[D.Theoryitem {text=str,
 | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 126 | name=SOME (nameparse "theoryitem" ty nameP toks), | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 127 | objtype=SOME ty}] | 
| 21867 | 128 | val plain_item = | 
| 129 | 		[D.Theoryitem {text=str,name=NONE,objtype=NONE}]
 | |
| 130 | val doccomment = | |
| 131 | 		[D.Doccomment {text=str}]
 | |
| 21637 | 132 | |
| 133 | val opt_overloaded = P.opt_keyword "overloaded"; | |
| 134 | ||
| 135 | val thmnameP = (* see isar_syn.ML/theoremsP *) | |
| 136 | let | |
| 22101 | 137 | val name_facts = P.and_list1 ((SpecParse.opt_thm_name "=" --| SpecParse.xthms1) >> #1) | 
| 138 | val theoremsP = P.opt_target |-- name_facts >> (fn [] => "" | x::_ => x) | |
| 21637 | 139 | in | 
| 140 | theoremsP | |
| 141 | end | |
| 142 | in | |
| 143 | (* TODO: ideally we would like to render terms properly, just as they are | |
| 144 | in output. This implies using PGML for symbols and variables/atoms. | |
| 145 | BUT it's rather tricky without having the correct concrete syntax to hand, | |
| 146 | and even if we did, we'd have to mess around here a whole lot more first | |
| 147 | to pick out the terms from the syntax. *) | |
| 148 | ||
| 149 | if member (op =) plain_items name then plain_item | |
| 150 | else case name of | |
| 21867 | 151 | "text" => doccomment | 
| 152 | | "text_raw" => doccomment | |
| 22405 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 153 | | "typedecl" => named_item (P.type_args |-- P.name) T.ObjType | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 154 | | "types" => named_item (P.type_args |-- P.name) T.ObjType | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 155 | | "classes" => named_item P.name I.ObjClass | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 156 | | "classrel" => named_item P.name I.ObjClass | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 157 | | "consts" => named_item (P.const >> #1) T.ObjTerm | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 158 | | "axioms" => named_item (SpecParse.spec_name >> (#1 o #1)) T.ObjTheorem | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 159 | | "defs" => named_item (opt_overloaded |-- SpecParse.spec_name >> (#1 o #1)) T.ObjTheorem | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 160 | | "constdefs" => named_item P.name (* FIXME: heavily simplified/wrong *) T.ObjTerm | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 161 | | "theorems" => named_item thmnameP I.ObjTheoremSet | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 162 | | "lemmas" => named_item thmnameP I.ObjTheoremSet | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 163 | | "oracle" => named_item P.name I.ObjOracle | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 164 | (* FIXME: locale needs to introduce a block *) | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 165 | | "locale" => named_item ((P.opt_keyword "open" >> not) |-- P.name) I.ObjLocale | 
| 21867 | 166 |                    | _ => (parse_warning ("Unrecognized thy-decl command: " ^ name); 
 | 
| 167 | plain_item) | |
| 21637 | 168 | end | 
| 169 | ||
| 170 | fun xmls_of_thy_goal (name,toks,str) = | |
| 171 | let | |
| 172 | (* see isar_syn.ML/gen_theorem *) | |
| 22101 | 173 | val statement = P.and_list1 (SpecParse.opt_thm_name ":" -- Scan.repeat1 P.propp); | 
| 21637 | 174 | val general_statement = | 
| 22101 | 175 | statement >> pair [] || Scan.repeat SpecParse.locale_element -- (P.$$$ "shows" |-- statement); | 
| 21637 | 176 | |
| 177 | val gen_theoremP = | |
| 22101 | 178 | P.opt_target -- Scan.optional (SpecParse.opt_thm_name ":" --| | 
| 179 |                  Scan.ahead (SpecParse.locale_keyword || P.$$$ "shows")) ("", []) --
 | |
| 21637 | 180 | general_statement >> (fn ((locale, a), (elems, concl)) => | 
| 181 | fst a) (* a : (bstring * Args.src list) *) | |
| 182 | ||
| 22101 | 183 | val nameP = P.opt_target |-- (Scan.optional ((SpecParse.opt_thm_name ":") >> #1) "") | 
| 22405 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 184 | val thmname = nameparse "opengoal" T.ObjTheorem nameP toks | 
| 21637 | 185 | in | 
| 21867 | 186 |             [D.Opengoal {thmname=SOME thmname, text=str},
 | 
| 22405 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 187 |              D.Openblock {metavarid=NONE,name=NONE,objtype=SOME I.ObjProofBody}]
 | 
| 21637 | 188 | end | 
| 189 | ||
| 21867 | 190 | fun xmls_of_qed (name,str) = | 
| 21637 | 191 | let val qedmarkup = | 
| 192 | (case name of | |
| 21867 | 193 |                      "sorry" => D.Postponegoal {text=str}
 | 
| 194 |                    | "oops"  => D.Giveupgoal {text=str}
 | |
| 195 |                    | "done"  => D.Closegoal {text=str}
 | |
| 196 |                    | "by"    => D.Closegoal {text=str}  (* nested or toplevel *)
 | |
| 197 |                    | "qed"   => D.Closegoal {text=str}  (* nested or toplevel *)
 | |
| 198 |                    | "."     => D.Closegoal {text=str}  (* nested or toplevel *)
 | |
| 199 |                    | ".."    => D.Closegoal {text=str}  (* nested or toplevel *)
 | |
| 21637 | 200 | | other => (* default to closegoal: may be wrong for extns *) | 
| 201 | (parse_warning | |
| 202 |                                        ("Unrecognized qed command: " ^ quote other);
 | |
| 21867 | 203 |                                        D.Closegoal {text=str}))
 | 
| 204 |         in qedmarkup :: [D.Closeblock {}] end
 | |
| 21637 | 205 | |
| 206 | fun xmls_of_kind kind (name,toks,str) = | |
| 207 | case kind of | |
| 21867 | 208 |       "control"        => [D.Badcmd {text=str}]
 | 
| 21971 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 209 |     | "diag"           => [D.Spuriouscmd {text=str}]
 | 
| 21637 | 210 | (* theory/files *) | 
| 211 | | "theory-begin" => xmls_of_thy_begin (name,toks,str) | |
| 212 | | "theory-decl" => xmls_of_thy_decl (name,toks,str) | |
| 22405 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 213 |     | "theory-heading" => [D.Closeblock {},
 | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 214 | 			   D.Doccomment {text=str},
 | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 215 | 			   D.Openblock {metavarid=NONE,name=NONE,objtype=SOME I.ObjTheoryBody}]
 | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 216 |     | "theory-script"  => [D.Theoryitem {name=NONE,objtype=SOME I.ObjTheoryDecl,text=str}]
 | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 217 |     | "theory-end"     => [D.Closeblock {}, D.Closetheory {text=str}]
 | 
| 21637 | 218 | (* proof control *) | 
| 219 | | "theory-goal" => xmls_of_thy_goal (name,toks,str) | |
| 21867 | 220 |     | "proof-heading"  => [D.Doccomment {text=str}]
 | 
| 22374 
db0b80b8371c
Add closeblock/openblock structure to proof-block commands
 aspinall parents: 
22160diff
changeset | 221 |     | "proof-goal"     => [D.Opengoal {text=str,thmname=NONE}, 
 | 
| 22405 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 222 | 			   D.Openblock {metavarid=NONE,name=NONE,objtype=SOME I.ObjProofBody}]
 | 
| 22374 
db0b80b8371c
Add closeblock/openblock structure to proof-block commands
 aspinall parents: 
22160diff
changeset | 223 |     | "proof-block"    => [D.Closeblock {}, 
 | 
| 
db0b80b8371c
Add closeblock/openblock structure to proof-block commands
 aspinall parents: 
22160diff
changeset | 224 | 			   D.Proofstep {text=str},
 | 
| 22405 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 225 | 			   D.Openblock {metavarid=NONE,name=NONE,objtype=SOME I.ObjProofBody}]
 | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 226 |     | "proof-open"     => [D.Openblock {metavarid=NONE,name=NONE,objtype=SOME I.ObjProofBody}, 
 | 
| 21867 | 227 | 			   D.Proofstep {text=str} ]
 | 
| 228 |     | "proof-close"    => [D.Proofstep {text=str}, D.Closeblock {}]
 | |
| 229 |     | "proof-script"   => [D.Proofstep {text=str}]
 | |
| 230 |     | "proof-chain"    => [D.Proofstep {text=str}]
 | |
| 231 |     | "proof-decl"     => [D.Proofstep {text=str}]
 | |
| 232 |     | "proof-asm"      => [D.Proofstep {text=str}]
 | |
| 22405 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 233 |     | "proof-asm-goal" => [D.Opengoal {text=str,thmname=NONE}, 
 | 
| 
7eef90be0db4
Add objtypes to openblock/closeblock and theory items (can be used to control folding).
 aspinall parents: 
22374diff
changeset | 234 | 			   D.Openblock {metavarid=NONE,name=NONE,objtype=NONE}]
 | 
| 21867 | 235 | | "qed" => xmls_of_qed (name,str) | 
| 236 | | "qed-block" => xmls_of_qed (name,str) | |
| 237 | | "qed-global" => xmls_of_qed (name,str) | |
| 238 | | other => (* default for anything else is "spuriouscmd", ***ignored for undo*** *) | |
| 21971 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 239 |       (parse_warning ("Internal inconsistency, unrecognized keyword kind: " ^ quote other ^
 | 
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 240 | "(treated as spuriouscmd)"); | 
| 21867 | 241 |        [D.Spuriouscmd {text=str}])
 | 
| 21637 | 242 | in | 
| 243 | ||
| 244 | fun xmls_of_transition (name,str,toks) = | |
| 245 | let | |
| 246 | val classification = Symtab.lookup (get_keywords_classification_table ()) name | |
| 247 | in case classification of | |
| 248 | SOME k => (xmls_of_kind k (name,toks,str)) | |
| 21971 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 249 | | NONE => (* an uncategorized keyword is treated as a theory item (maybe wrong) *) | 
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 250 |           (parse_warning ("Uncategorized command found: " ^ name ^ " -- using theoryitem");
 | 
| 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 251 |            [D.Theoryitem {name=NONE,objtype=NONE,text=str}])
 | 
| 21637 | 252 | end | 
| 253 | ||
| 22160 | 254 | |
| 255 | (* this would be a lot neater if we could match on toks! *) | |
| 21637 | 256 | fun markup_comment_whs [] = [] | 
| 22160 | 257 | | markup_comment_whs (toks as t::ts) = | 
| 21637 | 258 | let | 
| 259 | val (prewhs,rest) = take_prefix (OuterLex.is_kind OuterLex.Space) toks | |
| 260 | in | |
| 21902 | 261 | if not (null prewhs) then | 
| 21867 | 262 |             D.Whitespace {text=implode (map OuterLex.unparse prewhs)}
 | 
| 21637 | 263 | :: (markup_comment_whs rest) | 
| 264 | else | |
| 21867 | 265 |             D.Comment {text=OuterLex.unparse t} ::
 | 
| 21637 | 266 | (markup_comment_whs ts) | 
| 267 | end | |
| 268 | ||
| 269 | fun xmls_pre_cmd [] = ([],[]) | |
| 270 | | xmls_pre_cmd toks = | |
| 271 | let | |
| 272 | (* an extra token is needed to terminate the command *) | |
| 273 | val sync = OuterSyntax.scan "\\<^sync>"; | |
| 274 | val spaceP = Scan.bulk (Scan.one (fn t =>not (OuterLex.is_proper t)) >> OuterLex.val_of) >> implode | |
| 275 | val text_with_whs = | |
| 276 | ((spaceP || Scan.succeed "") -- | |
| 277 | (P.short_ident || P.long_ident || P.sym_ident || P.string || P.number || P.verbatim) >> op^) | |
| 278 | -- (spaceP || Scan.succeed "") >> op^ | |
| 279 | val (prewhs,rest1) = take_prefix (not o OuterLex.is_proper) (toks@sync) | |
| 21971 
513e1d82640d
Fix diagnostic kind (*are* spuriouscmd).  Add init to get right table for logic at startup.  Use theoryitem as default for unrecognised command.
 aspinall parents: 
21966diff
changeset | 280 | (* NB: this collapses doccomment,(doccomment|whitespace)* to a single doccomment *) | 
| 21637 | 281 | val (_,rest2) = (Scan.bulk (P.$$$ "--" |-- text_with_whs) >> implode || Scan.succeed "") rest1 | 
| 282 | val (postwhs,rest3) = take_prefix (not o OuterLex.is_proper) rest2 | |
| 283 | in | |
| 284 | ((markup_comment_whs prewhs) @ | |
| 285 | (if (length rest2 = length rest1) then [] | |
| 21867 | 286 | else | 
| 287 | [D.Doccomment | |
| 288 | 		   {text= implode
 | |
| 289 | (map OuterLex.unparse (Library.take (length rest1 - length rest2, rest1)))}] | |
| 290 | @ | |
| 291 | (markup_comment_whs postwhs)), | |
| 292 | Library.take (length rest3 - 1,rest3)) | |
| 21637 | 293 | end | 
| 294 | ||
| 295 | fun xmls_of_impropers toks = | |
| 296 | let | |
| 297 | val (xmls,rest) = xmls_pre_cmd toks | |
| 21867 | 298 | val unparsed = | 
| 299 | case rest of [] => [] | |
| 300 | | _ => [D.Unparseable | |
| 301 | 			      {text= implode (map OuterLex.unparse rest)}]
 | |
| 21637 | 302 | in | 
| 21867 | 303 | xmls @ unparsed | 
| 21637 | 304 | end | 
| 305 | ||
| 306 | end | |
| 307 | ||
| 308 | ||
| 309 | local | |
| 310 | (* we have to weave together the whitespace/comments with proper tokens | |
| 311 | to reconstruct the input. *) | |
| 312 | (* TODO: see if duplicating isar_output/parse_thy can help here *) | |
| 313 | ||
| 314 | fun match_tokens ([],us,vs) = (rev vs,us) (* used, unused *) | |
| 21902 | 315 | | match_tokens (t::ts, w::ws, vs) = | 
| 21966 | 316 | if (t: OuterLex.token) = w (* FIXME !? *) | 
| 21902 | 317 | then match_tokens (ts, ws, w::vs) | 
| 318 | else match_tokens (t::ts, ws, w::vs) | |
| 21637 | 319 |       | match_tokens _ = error ("match_tokens: mismatched input parse")
 | 
| 320 | in | |
| 21867 | 321 | fun pgip_parser str = | 
| 21637 | 322 | let | 
| 323 | (* parsing: See outer_syntax.ML/toplevel_source *) | |
| 324 | fun parse_loop ([],[],xmls) = xmls | |
| 325 | | parse_loop ([],itoks,xmls) = xmls @ (xmls_of_impropers itoks) | |
| 326 | | parse_loop ((nm,toks,_)::trs,itoks,xmls) = | |
| 327 | let | |
| 328 | (* first proper token after whitespace/litcomment/whitespace is command *) | |
| 329 | val (xmls_pre_cmd,cmditoks') = xmls_pre_cmd itoks | |
| 22160 | 330 | val (cmdtok,itoks') = | 
| 331 | case cmditoks' of x::xs => (x,xs) | |
| 332 |                                    | _ => error("proof_general/parse_loop impossible")
 | |
| 21637 | 333 | val (utoks,itoks'') = match_tokens (toks,itoks',[]) | 
| 334 | (* FIXME: should take trailing [w/s+] semicolon too in utoks *) | |
| 335 | ||
| 336 | val str = implode (map OuterLex.unparse (cmdtok::utoks)) | |
| 337 | ||
| 338 | val xmls_tr = xmls_of_transition (nm,str,toks) | |
| 339 | in | |
| 340 | parse_loop(trs,itoks'',xmls @ xmls_pre_cmd @ xmls_tr) | |
| 341 | end | |
| 342 | in | |
| 343 | (let val toks = OuterSyntax.scan str | |
| 344 | in | |
| 345 | parse_loop (OuterSyntax.read toks, toks, []) | |
| 346 | end) | |
| 21867 | 347 |            handle _ => [D.Unparseable {text=str}]
 | 
| 21637 | 348 | end | 
| 349 | end | |
| 350 | ||
| 351 | end |