| author | wenzelm | 
| Sun, 15 Nov 2009 15:14:02 +0100 | |
| changeset 33696 | 2c7c79ca6c23 | 
| parent 33038 | 8f9594c31de4 | 
| child 36315 | e859879079c8 | 
| permissions | -rw-r--r-- | 
| 23797 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 1 | (* Title: Pure/ProofGeneral/pgip_parser.ML | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 2 | Author: David Aspinall and Makarius | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 3 | |
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 4 | Parsing theory sources without execution (via keyword classification). | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 5 | *) | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 6 | |
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 7 | signature PGIP_PARSER = | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 8 | sig | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 9 | val pgip_parser: Position.T -> string -> PgipMarkup.pgipdocument | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 10 | end | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 11 | |
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 12 | structure PgipParser: PGIP_PARSER = | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 13 | struct | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 14 | |
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 15 | structure K = OuterKeyword; | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 16 | structure D = PgipMarkup; | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 17 | structure I = PgipIsabelle; | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 18 | |
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 19 | |
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 20 | fun badcmd text = [D.Badcmd {text = text}];
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 21 | |
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 22 | fun thy_begin text = | 
| 32466 | 23 | (case try (Thy_Header.read Position.none) (Source.of_string text) of | 
| 24192 
4eccd4bb8b64
PGIP change: thyname is optional in opentheory, markup even in case of header parse failure
 aspinall parents: 
23868diff
changeset | 24 |     NONE => D.Opentheory {thyname = NONE, parentnames = [], text = text}
 | 
| 23797 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 25 | | SOME (name, imports, _) => | 
| 24192 
4eccd4bb8b64
PGIP change: thyname is optional in opentheory, markup even in case of header parse failure
 aspinall parents: 
23868diff
changeset | 26 |        D.Opentheory {thyname = SOME name, parentnames = imports, text = text})
 | 
| 
4eccd4bb8b64
PGIP change: thyname is optional in opentheory, markup even in case of header parse failure
 aspinall parents: 
23868diff
changeset | 27 |   :: [D.Openblock {metavarid = NONE, name = NONE, objtype = SOME I.ObjTheoryBody}];
 | 
| 23797 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 28 | |
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 29 | fun thy_heading text = | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 30 |   [D.Closeblock {},
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 31 |    D.Doccomment {text = text},
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 32 |    D.Openblock {metavarid = NONE, name = NONE, objtype = SOME I.ObjTheoryBody}];
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 33 | |
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 34 | fun thy_decl text = | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 35 |   [D.Theoryitem {name = NONE, objtype = SOME I.ObjTheoryDecl, text = text}];
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 36 | |
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 37 | fun goal text = | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 38 |   [D.Opengoal {thmname = NONE, text = text},
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 39 |    D.Openblock {metavarid = NONE, name = NONE, objtype = SOME I.ObjProofBody}];
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 40 | |
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 41 | fun prf_block text = | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 42 |   [D.Closeblock {},
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 43 |    D.Proofstep {text = text},
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 44 |    D.Openblock {metavarid = NONE, name = NONE, objtype = SOME I.ObjProofBody}];
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 45 | |
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 46 | fun prf_open text = | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 47 |  [D.Openblock {metavarid = NONE, name = NONE, objtype = SOME I.ObjProofBody},
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 48 |   D.Proofstep {text = text}];
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 49 | |
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 50 | fun proofstep text = [D.Proofstep {text = text}];
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 51 | fun closegoal text = [D.Closegoal {text = text}, D.Closeblock {}];
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 52 | |
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 53 | |
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 54 | fun command k f = Symtab.update_new (OuterKeyword.kind_of k, f); | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 55 | |
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 56 | val command_keywords = Symtab.empty | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 57 | |> command K.control badcmd | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 58 |   |> command K.diag             (fn text => [D.Spuriouscmd {text = text}])
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 59 | |> command K.thy_begin thy_begin | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 60 | |> command K.thy_switch badcmd | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 61 |   |> command K.thy_end          (fn text => [D.Closeblock {}, D.Closetheory {text = text}])
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 62 | |> command K.thy_heading thy_heading | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 63 | |> command K.thy_decl thy_decl | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 64 | |> command K.thy_script thy_decl | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 65 | |> command K.thy_goal goal | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 66 | |> command K.qed closegoal | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 67 | |> command K.qed_block closegoal | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 68 |   |> command K.qed_global       (fn text => [D.Giveupgoal {text = text}])
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 69 |   |> command K.prf_heading      (fn text => [D.Doccomment {text = text}])
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 70 | |> command K.prf_goal goal | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 71 | |> command K.prf_block prf_block | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 72 | |> command K.prf_open prf_open | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 73 |   |> command K.prf_close        (fn text => [D.Proofstep {text = text}, D.Closeblock {}])
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 74 | |> command K.prf_chain proofstep | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 75 | |> command K.prf_decl proofstep | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 76 | |> command K.prf_asm proofstep | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 77 | |> command K.prf_asm_goal goal | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 78 | |> command K.prf_script proofstep; | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 79 | |
| 33038 | 80 | val _ = subset (op =) (OuterKeyword.kinds, Symtab.keys command_keywords) | 
| 23797 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 81 | orelse sys_error "Incomplete coverage of command keywords"; | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 82 | |
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 83 | fun parse_command "sorry" text = [D.Postponegoal {text = text}, D.Closeblock {}]
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 84 | | parse_command name text = | 
| 27353 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 wenzelm parents: 
24192diff
changeset | 85 | (case OuterKeyword.command_keyword name of | 
| 23797 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 86 |         NONE => [D.Unparseable {text = text}]
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 87 | | SOME k => | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 88 | (case Symtab.lookup command_keywords (OuterKeyword.kind_of k) of | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 89 |             NONE => [D.Unparseable {text = text}]
 | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 90 | | SOME f => f text)); | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 91 | |
| 27841 | 92 | fun parse_span span = | 
| 93 | let | |
| 29315 
b074c05f00ad
renamed ThyEdit (in thy_edit.ML) to ThySyntax (in thy_syntax.ML);
 wenzelm parents: 
28035diff
changeset | 94 | val kind = ThySyntax.span_kind span; | 
| 
b074c05f00ad
renamed ThyEdit (in thy_edit.ML) to ThySyntax (in thy_syntax.ML);
 wenzelm parents: 
28035diff
changeset | 95 | val toks = ThySyntax.span_content span; | 
| 
b074c05f00ad
renamed ThyEdit (in thy_edit.ML) to ThySyntax (in thy_syntax.ML);
 wenzelm parents: 
28035diff
changeset | 96 | val text = implode (map (PrintMode.setmp [] ThySyntax.present_token) toks); | 
| 27841 | 97 | in | 
| 23797 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 98 | (case kind of | 
| 29315 
b074c05f00ad
renamed ThyEdit (in thy_edit.ML) to ThySyntax (in thy_syntax.ML);
 wenzelm parents: 
28035diff
changeset | 99 | ThySyntax.Command name => parse_command name text | 
| 
b074c05f00ad
renamed ThyEdit (in thy_edit.ML) to ThySyntax (in thy_syntax.ML);
 wenzelm parents: 
28035diff
changeset | 100 |     | ThySyntax.Ignored => [D.Whitespace {text = text}]
 | 
| 
b074c05f00ad
renamed ThyEdit (in thy_edit.ML) to ThySyntax (in thy_syntax.ML);
 wenzelm parents: 
28035diff
changeset | 101 |     | ThySyntax.Malformed => [D.Unparseable {text = text}])
 | 
| 23797 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 102 | end; | 
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 103 | |
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 104 | |
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 105 | fun pgip_parser pos str = | 
| 29315 
b074c05f00ad
renamed ThyEdit (in thy_edit.ML) to ThySyntax (in thy_syntax.ML);
 wenzelm parents: 
28035diff
changeset | 106 | maps parse_span (ThySyntax.parse_spans (OuterKeyword.get_lexicons ()) pos str); | 
| 23797 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 107 | |
| 
f4dbbffbfe06
Parsing theory sources without execution (via keyword classification).
 wenzelm parents: diff
changeset | 108 | end; |