| author | Christian Urban <urbanc@in.tum.de> | 
| Tue, 16 Dec 2008 08:08:44 +0000 | |
| changeset 56 | 126646f2aa88 | 
| parent 54 | 1783211b3494 | 
| child 58 | f3794c231898 | 
| permissions | -rw-r--r-- | 
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1 | theory Parsing | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 2 | imports Base | 
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 3 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 4 | begin | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 5 | |
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 6 | |
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 7 | chapter {* Parsing *}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 8 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 9 | text {*
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 10 | |
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 11 |   Isabelle distinguishes between \emph{outer} and \emph{inner} syntax. 
 | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 12 |   Theory commands, such as \isacommand{definition}, \isacommand{inductive} and so
 | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 13 | on, belong to the outer syntax, whereas items inside double quotation marks, such | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 14 | as terms, types and so on, belong to the inner syntax. For parsing inner syntax, | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 15 | Isabelle uses a rather general and sophisticated algorithm due to Earley, which | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 16 | is driven by priority grammars. Parsers for outer syntax are built up by functional | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 17 | parsing combinators. These combinators are a well-established technique for parsing, | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 18 |   which has, for example, been described in Paulson's classic ML-book \cite{paulson-ml2}.
 | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 19 | Isabelle developers are usually concerned with writing these outer syntax parsers, | 
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 20 | either for new definitional packages or for calling tactics with specific arguments. | 
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 21 | |
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 22 |   \begin{readmore}
 | 
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 23 | The library | 
| 50 | 24 | for writing parser combinators is split up, roughly, into two parts. | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 25 | The first part consists of a collection of generic parser combinators defined | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 26 |   in the structure @{ML_struct Scan} in the file 
 | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 27 |   @{ML_file "Pure/General/scan.ML"}. The second part of the library consists of 
 | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 28 | combinators for dealing with specific token types, which are defined in the | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 29 |   structure @{ML_struct OuterParse} in the file 
 | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 30 |   @{ML_file "Pure/Isar/outer_parse.ML"}.
 | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 31 |   \end{readmore}
 | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 32 | |
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 33 | *} | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 34 | |
| 49 | 35 | section {* Building Generic Parsers *}
 | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 36 | |
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 37 | text {*
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 38 | |
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 39 | Let us first have a look at parsing strings using generic parsing combinators. | 
| 50 | 40 |   The function @{ML "(op $$)"} takes a string as argument and will ``consume'' this string from 
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 41 | a given input list of strings. ``Consume'' in this context means that it will | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 42 | return a pair consisting of this string and the rest of the input list. | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 43 | For example: | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 44 | |
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 45 |   @{ML_response [display] "($$ \"h\") (explode \"hello\")" "(\"h\", [\"e\", \"l\", \"l\", \"o\"])"}
 | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 46 |   @{ML_response [display] "($$ \"w\") (explode \"world\")" "(\"w\", [\"o\", \"r\", \"l\", \"d\"])"}
 | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 47 | |
| 52 | 48 | This function will either succeed (as in the two examples above) or raise the exception | 
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 49 |   @{ML_text "FAIL"} if no string can be consumed. For example trying to parse
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 50 | |
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 51 |   @{ML_response_fake [display] "($$ \"x\") (explode \"world\")" 
 | 
| 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 52 | "Exception FAIL raised"} | 
| 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 53 | |
| 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 54 |   will raise the exception @{ML_text "FAIL"}.
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 55 | There are three exceptions used in the parsing combinators: | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 56 | |
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 57 |   \begin{itemize}
 | 
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 58 |   \item @{ML_text "FAIL"} is used to indicate that alternative routes of parsing 
 | 
| 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 59 | might be explored. | 
| 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 60 |   \item @{ML_text "MORE"} indicates that there is not enough input for the parser. For example 
 | 
| 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 61 |   in @{ML_text "($$ \"h\") []"}.
 | 
| 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 62 |   \item @{ML_text "ABORT"} is the exception which is raised when a dead end is reached. 
 | 
| 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 63 |   It is used for example in the function @{ML "(op !!)"} (see below).
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 64 |   \end{itemize}
 | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 65 | |
| 50 | 66 | However, note that these exceptions are private to the parser and cannot be accessed | 
| 49 | 67 | by the programmer (for example to handle them). | 
| 68 | ||
| 69 |   Slightly more general than the parser @{ML "(op $$)"} is the function @{ML
 | |
| 70 | Scan.one}, in that it takes a predicate as argument and then parses exactly | |
| 52 | 71 | one item from the input list satisfying this predicate. For example the | 
| 49 | 72 |   following parser either consumes an @{ML_text [quotes] "h"} or a @{ML_text
 | 
| 73 | [quotes] "w"}: | |
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 74 | |
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 75 | |
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 76 | @{ML_response [display] 
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 77 | "let | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 78 | val hw = Scan.one (fn x => x = \"h\" orelse x = \"w\") | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 79 | val input1 = (explode \"hello\") | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 80 | val input2 = (explode \"world\") | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 81 | in | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 82 | (hw input1, hw input2) | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 83 | end" | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 84 | "((\"h\", [\"e\", \"l\", \"l\", \"o\"]),(\"w\", [\"o\", \"r\", \"l\", \"d\"]))"} | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 85 | |
| 52 | 86 |   Two parser can be connected in sequence by using the function @{ML "(op --)"}. 
 | 
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 87 |   For example parsing @{ML_text "h"}, @{ML_text "e"} and @{ML_text "l"} in this 
 | 
| 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 88 | sequence can be achieved by | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 89 | |
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 90 |   @{ML_response [display] "(($$ \"h\") -- ($$ \"e\") -- ($$ \"l\")) (explode \"hello\")"
 | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 91 | "(((\"h\", \"e\"), \"l\"), [\"l\", \"o\"])"} | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 92 | |
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 93 | Note how the result of consumed strings builds up on the left as nested pairs. | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 94 | |
| 56 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 95 | If, as in the previous example, one wants to parse a particular string, | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 96 |   then one should rather use the function @{ML Scan.this_string}:
 | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 97 | |
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 98 |   @{ML_response [display] "Scan.this_string \"hel\" (explode \"hello\")"
 | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 99 | "(\"hel\", [\"l\", \"o\"])"} | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 100 | |
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 101 |   Parsers that explore alternatives can be constructed using the function @{ML
 | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 102 |   "(op ||)"}. For example, the parser @{ML "(p || q)" for p q} returns the
 | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 103 |   result of @{ML_text "p"}, in case it succeeds, otherwise it returns the
 | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 104 |   result of @{ML_text "q"}. For example
 | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 105 | |
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 106 | |
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 107 | @{ML_response [display] 
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 108 | "let | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 109 | val hw = ($$ \"h\") || ($$ \"w\") | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 110 | val input1 = (explode \"hello\") | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 111 | val input2 = (explode \"world\") | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 112 | in | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 113 | (hw input1, hw input2) | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 114 | end" | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 115 | "((\"h\", [\"e\", \"l\", \"l\", \"o\"]), (\"w\", [\"o\", \"r\", \"l\", \"d\"]))"} | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 116 | |
| 52 | 117 |   The functions @{ML "(op |--)"} and @{ML "(op --|)"} work like the sequencing function 
 | 
| 50 | 118 | for parsers, except that they discard the item being parsed by the first (respectively second) | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 119 | parser. For example | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 120 | |
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 121 | @{ML_response [display]
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 122 | "let | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 123 | val just_e = ($$ \"h\") |-- ($$ \"e\") | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 124 | val just_h = ($$ \"h\") --| ($$ \"e\") | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 125 | val input = (explode \"hello\") | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 126 | in | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 127 | (just_e input, just_h input) | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 128 | end" | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 129 | "((\"e\", [\"l\", \"l\", \"o\"]),(\"h\", [\"l\", \"l\", \"o\"]))"} | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 130 | |
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 131 |   The parser @{ML "Scan.optional p x" for p x} returns the result of the parser 
 | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 132 |   @{ML_text "p"}, if it succeeds; otherwise it returns 
 | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 133 |   the default value @{ML_text "x"}. For example
 | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 134 | |
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 135 | @{ML_response [display]
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 136 | "let | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 137 | val p = Scan.optional ($$ \"h\") \"x\" | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 138 | val input1 = (explode \"hello\") | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 139 | val input2 = (explode \"world\") | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 140 | in | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 141 | (p input1, p input2) | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 142 | end" | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 143 | "((\"h\", [\"e\", \"l\", \"l\", \"o\"]), (\"x\", [\"w\", \"o\", \"r\", \"l\", \"d\"]))"} | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 144 | |
| 49 | 145 |   The function @{ML Scan.option} works similarly, except no default value can
 | 
| 50 | 146 |   be given. Instead, the result is wrapped as an @{text "option"}-type. For example:
 | 
| 147 | ||
| 148 | @{ML_response [display]
 | |
| 149 | "let | |
| 150 | val p = Scan.option ($$ \"h\") | |
| 151 | val input1 = (explode \"hello\") | |
| 152 | val input2 = (explode \"world\") | |
| 153 | in | |
| 154 | (p input1, p input2) | |
| 155 | end" "((SOME \"h\", [\"e\", \"l\", \"l\", \"o\"]), (NONE, [\"w\", \"o\", \"r\", \"l\", \"d\"]))"} | |
| 49 | 156 | |
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 157 |   The function @{ML "(op !!)"} helps to produce appropriate error messages
 | 
| 43 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 158 |   during parsing. For example if one wants to parse that @{ML_text p} is immediately 
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 159 |   followed by @{ML_text q}, or start a completely different parser @{ML_text r},
 | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 160 | one might write | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 161 | |
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 162 |   @{ML [display] "(p -- q) || r" for p q r}
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 163 | |
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 164 | However, this parser is problematic for producing an appropriate error message, in case | 
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 165 |   the parsing of @{ML "(p -- q)" for p q} fails. Because in that case one loses with the parser
 | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 166 | above the information | 
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 167 |   that @{ML_text p} should be followed by @{ML_text q}. To see this consider the case in
 | 
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 168 |   which @{ML_text p} 
 | 
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 169 |   is present in the input, but not @{ML_text q}. That means @{ML "(p -- q)" for p q} will fail 
 | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 170 | and the | 
| 52 | 171 |   alternative parser @{ML_text r} will be tried. However in many circumstance this will be the wrong
 | 
| 43 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 172 | parser for the input ``p-followed-by-q'' and therefore will also fail. The error message is then | 
| 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 173 | caused by the | 
| 52 | 174 |   failure of @{ML_text r}, not by the absence of @{ML_text q} in the input. This kind of situation
 | 
| 175 |   can be avoided by using the function @{ML "(op !!)"}. This function aborts the whole process of
 | |
| 50 | 176 | parsing in case of a failure and invokes an error message. For example if we invoke the parser | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 177 | |
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 178 |   @{ML [display] "(!! (fn _ => \"foo\") ($$ \"h\"))"}
 | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 179 | |
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 180 |   on @{ML_text [quotes] "hello"}, the parsing succeeds
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 181 | |
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 182 |   @{ML_response [display] 
 | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 183 | "(!! (fn _ => \"foo\") ($$ \"h\")) (explode \"hello\")" | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 184 | "(\"h\", [\"e\", \"l\", \"l\", \"o\"])"} | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 185 | |
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 186 |   but if we invoke it on @{ML_text [quotes] "world"}
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 187 | |
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 188 |   @{ML_response_fake [display] "(!! (fn _ => \"foo\") ($$ \"h\")) (explode \"world\")"
 | 
| 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 189 | "Exception ABORT raised"} | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 190 | |
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 191 |   the parsing aborts and the error message @{ML_text "foo"} is printed out. In order to
 | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 192 | see the error message properly, we need to prefix the parser with the function | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 193 |   @{ML "Scan.error"}. For example
 | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 194 | |
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 195 |   @{ML_response_fake [display] "Scan.error ((!! (fn _ => \"foo\") ($$ \"h\")))"
 | 
| 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 196 | "Exception Error \"foo\" raised"} | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 197 | |
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 198 |   This ``prefixing'' is usually done by wrappers such as @{ML "OuterSyntax.command"} 
 | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 199 | (FIXME: give reference to later place). | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 200 | |
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 201 |   Returning to our example of parsing @{ML "(p -- q) || r" for p q r}. If we want
 | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 202 | to generate the correct error message for p-followed-by-q, then | 
| 50 | 203 | we have to write: | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 204 | *} | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 205 | |
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 206 | ML {* 
 | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 207 | fun p_followed_by_q p q r = | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 208 | let | 
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 209 | val err = (fn _ => p ^ " is not followed by " ^ q) | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 210 | in | 
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 211 | (($$ p) -- (!! err ($$ q))) || (($$ r) -- ($$ r)) | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 212 | end | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 213 | *} | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 214 | |
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 215 | |
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 216 | text {*
 | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 217 | Running this parser with | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 218 | |
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 219 |   @{ML_response_fake [display] "Scan.error (p_followed_by_q \"h\" \"e\" \"w\") (explode \"holle\")"
 | 
| 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 220 | "Exception ERROR \"h is not followed by e\" raised"} | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 221 | |
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 222 | gives the correct error message. Running it with | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 223 | |
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 224 |   @{ML_response [display] "Scan.error (p_followed_by_q \"h\" \"e\" \"w\") (explode \"wworld\")"
 | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 225 | "((\"w\", \"w\"), [\"o\", \"r\", \"l\", \"d\"])"} | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 226 | |
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 227 | yields the expected parsing. | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 228 | |
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 229 |   The function @{ML "Scan.repeat p" for p} will apply a parser @{ML_text p} as 
 | 
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 230 | often as it succeeds. For example | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 231 | |
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 232 |   @{ML_response [display] "Scan.repeat ($$ \"h\") (explode \"hhhhello\")" 
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 233 | "([\"h\", \"h\", \"h\", \"h\"], [\"e\", \"l\", \"l\", \"o\"])"} | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 234 | |
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 235 |   Note that @{ML "Scan.repeat"} stores the parsed items in a list. The function
 | 
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 236 |   @{ML "Scan.repeat1"} is similar, but requires that the parser @{ML_text "p"} 
 | 
| 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 237 | succeeds at least once. | 
| 48 
609f9ef73494
fixed FIXME's in fake responses
 Christian Urban <urbanc@in.tum.de> parents: 
47diff
changeset | 238 | |
| 49 | 239 |   Also note that the parser would have aborted with the exception @{ML_text MORE}, if
 | 
| 50 | 240 |   we had run it only on just @{ML_text [quotes] "hhhh"}. This can be avoided using
 | 
| 49 | 241 |   the wrapper @{ML Scan.finite} and the ``stopper-token'' @{ML Symbol.stopper}. With
 | 
| 242 | them we can write | |
| 243 | ||
| 244 |   @{ML_response [display] "Scan.finite Symbol.stopper (Scan.repeat ($$ \"h\")) (explode \"hhhh\")" 
 | |
| 245 | "([\"h\", \"h\", \"h\", \"h\"], [])"} | |
| 246 | ||
| 50 | 247 | However, this kind of manually wrapping needs to be done only very rarely | 
| 248 | in practise, because it is already done by the infrastructure for you. | |
| 49 | 249 | |
| 56 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 250 |   The function @{ML Scan.repeat} can be used with @{ML Scan.one} to read any 
 | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 251 | string as in | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 252 | |
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 253 |   @{ML_response [display] 
 | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 254 | "let | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 255 | val p = Scan.repeat (Scan.one Symbol.not_eof) | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 256 | val input = (explode \"foo bar foo\") | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 257 | in | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 258 | Scan.finite Symbol.stopper p input | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 259 | end" | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 260 | "([\"f\", \"o\", \"o\", \" \", \"b\", \"a\", \"r\", \" \", \"f\", \"o\", \"o\"], [])"} | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 261 | |
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 262 |   where the function @{ML Symbol.not_eof} ensures that we do not read beyond the 
 | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 263 | end of the input string. | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 264 | |
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 265 |   The function @{ML "Scan.unless p q" for p q} takes two parsers: if the first one can 
 | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 266 | parse the input then the whole parser fails; if not, then the second is tried. Therefore | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 267 | |
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 268 |   @{ML_response_fake_both [display] "Scan.unless ($$ \"h\") ($$ \"w\") (explode \"hello\")"
 | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 269 | "Exception FAIL raised"} | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 270 | |
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 271 | fails, while | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 272 | |
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 273 |   @{ML_response [display] "Scan.unless ($$ \"h\") ($$ \"w\") (explode \"world\")"
 | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 274 | "(\"w\",[\"o\", \"r\", \"l\", \"d\"])"} | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 275 | |
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 276 | succeeds. | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 277 | |
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 278 |   The functions @{ML Scan.repeat} and @{ML Scan.unless} can be combined to read any
 | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 279 | input until a certain marker symbol is reached. In the example below the marker | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 280 |   symbol is a @{text [quotes] "*"}:
 | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 281 | |
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 282 |   @{ML_response [display]
 | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 283 | "let | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 284 | val p = Scan.repeat (Scan.unless ($$ \"*\") (Scan.one Symbol.not_eof)) | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 285 | val input1 = (explode \"fooooo\") | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 286 | val input2 = (explode \"foo*ooo\") | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 287 | in | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 288 | (Scan.finite Symbol.stopper p input1, | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 289 | Scan.finite Symbol.stopper p input2) | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 290 | end" | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 291 | "(([\"f\", \"o\", \"o\", \"o\", \"o\", \"o\"], []), | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 292 | ([\"f\", \"o\", \"o\"], [\"*\", \"o\", \"o\", \"o\"]))"} | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 293 | |
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 294 | After parsing succeeded, one nearly always wants to apply a function on the parsed | 
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 295 |   items. This is done using the function @{ML "(p >> f)" for p f} which runs 
 | 
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 296 |   first the parser @{ML_text p} and upon successful completion applies the 
 | 
| 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 297 |   function @{ML_text f} to the result. For example
 | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 298 | |
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 299 | @{ML_response [display]
 | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 300 | "let | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 301 | fun double (x,y) = (x^x,y^y) | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 302 | in | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 303 | (($$ \"h\") -- ($$ \"e\") >> double) (explode \"hello\") | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 304 | end" | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 305 | "((\"hh\", \"ee\"), [\"l\", \"l\", \"o\"])"} | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 306 | |
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 307 | doubles the two parsed input strings. | 
| 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 308 | |
| 56 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 309 |   \begin{exercise}\label{ex:scancmts}
 | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 310 | Write a parser parses an input string so that any comment enclosed inside | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 311 |   @{text "(*\<dots>*)"} is enclosed inside @{text "(**\<dots>**)"} in the output
 | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 312 |   string. To enclose a string, you can use the function @{ML "enclose s1 s2 s"
 | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 313 |   for s1 s2 s} which produces the string @{ML "s1^s^s2" for s1 s2 s}.
 | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 314 |   \end{exercise}
 | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 315 | |
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 316 | |
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 317 |   The function @{ML Scan.lift} takes a parser and a pair as arguments. This function applies
 | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 318 | the given parser to the second component of the pair and leaves the first component | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 319 | untouched. For example | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 320 | |
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 321 | @{ML_response [display]
 | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 322 | "Scan.lift (($$ \"h\") -- ($$ \"e\")) (1,(explode \"hello\"))" | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 323 | "((\"h\", \"e\"), (1, [\"l\", \"l\", \"o\"]))"} | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 324 | |
| 43 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 325 | (FIXME: In which situations is this useful? Give examples.) | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 326 | *} | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 327 | |
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 328 | section {* Parsing Theory Syntax *}
 | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 329 | |
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 330 | text {*
 | 
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 331 | Most of the time, however, Isabelle developers have to deal with parsing tokens, not strings. | 
| 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 332 | This is because the parsers for the theory syntax, as well as the parsers for the | 
| 43 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 333 | argument syntax of proof methods and attributes use the type | 
| 44 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 334 |   @{ML_type OuterLex.token} (which is identical to the type @{ML_type OuterParse.token}).
 | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 335 | |
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 336 |   \begin{readmore}
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 337 | The parser functions for the theory syntax are contained in the structure | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 338 |   @{ML_struct OuterParse} defined in the file @{ML_file  "Pure/Isar/outer_parse.ML"}.
 | 
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 339 |   The definition for tokens is in the file @{ML_file "Pure/Isar/outer_lex.ML"}.
 | 
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 340 |   \end{readmore}
 | 
| 44 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 341 | |
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 342 |   The structure @{ML_struct OuterLex} defines several kinds of token (for example 
 | 
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 343 |   @{ML "Ident" in OuterLex} for identifiers, @{ML "Keyword" in OuterLex} for keywords and
 | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 344 |   @{ML "Command" in OuterLex} for commands). Some token parsers take into account the 
 | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 345 | kind of token. | 
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 346 | *} | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 347 | |
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 348 | text {*
 | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 349 | For the examples below, we can generate a token list out of a string using | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 350 |   the function @{ML "OuterSyntax.scan"}, which we give below @{ML
 | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 351 | "Position.none"} as argument since, at the moment, we are not interested in | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 352 | generating precise error messages. The following | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 353 | |
| 44 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 354 | |
| 49 | 355 | @{ML_response_fake [display] "OuterSyntax.scan Position.none \"hello world\"" 
 | 
| 50 | 356 | "[Token (\<dots>,(Ident, \"hello\"),\<dots>), | 
| 357 | Token (\<dots>,(Space, \" \"),\<dots>), | |
| 358 | Token (\<dots>,(Ident, \"world\"),\<dots>)]"} | |
| 359 | ||
| 360 | produces three tokens where the first and the last are identifiers, since | |
| 361 |   @{ML_text [quotes] "hello"} and @{ML_text [quotes] "world"} do not match any
 | |
| 362 |   other syntactic category.\footnote{Note that because of a possible a bug in
 | |
| 363 |   the PolyML runtime system the result is printed as @{text "?"}, instead of
 | |
| 364 | the token.} The second indicates a space. | |
| 365 | ||
| 366 | Many parsing functions later on will require spaces, comments and the like | |
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 367 | to have already been filtered out. So from now on we are going to use the | 
| 50 | 368 |   functions @{ML filter} and @{ML OuterLex.is_proper} do this. For example
 | 
| 369 | ||
| 44 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 370 | |
| 50 | 371 | @{ML_response_fake [display]
 | 
| 372 | "let | |
| 373 | val input = OuterSyntax.scan Position.none \"hello world\" | |
| 374 | in | |
| 375 | filter OuterLex.is_proper input | |
| 376 | end" | |
| 377 | "[Token (\<dots>,(Ident, \"hello\"), \<dots>), Token (\<dots>,(Ident, \"world\"), \<dots>)]"} | |
| 378 | ||
| 379 | For convenience we are going to use the function | |
| 380 | ||
| 381 | *} | |
| 382 | ||
| 383 | ML {* 
 | |
| 384 | fun filtered_input str = | |
| 385 | filter OuterLex.is_proper (OuterSyntax.scan Position.none str) | |
| 386 | *} | |
| 387 | ||
| 388 | text {*
 | |
| 389 | ||
| 48 
609f9ef73494
fixed FIXME's in fake responses
 Christian Urban <urbanc@in.tum.de> parents: 
47diff
changeset | 390 | If we parse | 
| 44 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 391 | |
| 50 | 392 | @{ML_response_fake [display] 
 | 
| 393 | "filtered_input \"inductive | for\"" | |
| 394 | "[Token (\<dots>,(Command, \"inductive\"),\<dots>), | |
| 395 | Token (\<dots>,(Keyword, \"|\"),\<dots>), | |
| 396 | Token (\<dots>,(Keyword, \"for\"),\<dots>)]"} | |
| 44 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 397 | |
| 52 | 398 | we obtain a list consisting of only a command and two keyword tokens. | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 399 | If you want to see which keywords and commands are currently known, use | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 400 |   the following (you might have to adjust the @{ML print_depth} in order to
 | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 401 | see the complete list): | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 402 | |
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 403 | @{ML_response_fake [display] 
 | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 404 | "let | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 405 | val (keywords, commands) = OuterKeyword.get_lexicons () | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 406 | in | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 407 | (Scan.dest_lexicon commands, Scan.dest_lexicon keywords) | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 408 | end" | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 409 | "([\"}\",\"{\",\<dots>],[\"\<rightleftharpoons>\",\"\<leftharpoondown>\",\<dots>])"}
 | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 410 | |
| 44 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 411 |   Now the parser @{ML "OuterParse.$$$"} parses a single keyword. For example
 | 
| 50 | 412 | |
| 44 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 413 | @{ML_response [display]
 | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 414 | "let | 
| 50 | 415 | val input1 = filtered_input \"where for\" | 
| 416 | val input2 = filtered_input \"| in\" | |
| 44 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 417 | in | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 418 | (OuterParse.$$$ \"where\" input1, OuterParse.$$$ \"|\" input2) | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 419 | end" | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 420 | "((\"where\",\<dots>),(\"|\",\<dots>))"} | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 421 | |
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 422 |   Like before, we can sequentially connect parsers with @{ML "(op --)"}. For example 
 | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 423 | |
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 424 | @{ML_response [display]
 | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 425 | "let | 
| 50 | 426 | val input = filtered_input \"| in\" | 
| 44 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 427 | in | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 428 | (OuterParse.$$$ \"|\" -- OuterParse.$$$ \"in\") input | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 429 | end" | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 430 | "((\"|\",\"in\"),[])"} | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 431 | |
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 432 |   The parser @{ML "OuterParse.enum s p" for s p} parses a possibly empty 
 | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 433 |   list of items recognised by the parser @{ML_text p}, where the items being parsed
 | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 434 |   are separated by the string @{ML_text s}. For example
 | 
| 44 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 435 | |
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 436 | @{ML_response [display]
 | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 437 | "let | 
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 438 | val input = filtered_input \"in | in | in foo\" | 
| 44 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 439 | in | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 440 | (OuterParse.enum \"|\" (OuterParse.$$$ \"in\")) input | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 441 | end" | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 442 | "([\"in\",\"in\",\"in\"],[\<dots>])"} | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 443 | |
| 50 | 444 |   @{ML "OuterParse.enum1"} works similarly, except that the parsed list must
 | 
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 445 |   be non-empty. Note that we had to add an @{ML_text [quotes] "foo"} at the end
 | 
| 50 | 446 | of the parsed string, otherwise the parser would have consumed all tokens | 
| 447 |   and then failed with the exception @{ML_text "MORE"}. Like in the previous
 | |
| 448 |   section, we can avoid this exception using the wrapper @{ML
 | |
| 449 |   Scan.finite}. This time, however, we have to use the ``stopper-token'' @{ML
 | |
| 450 | OuterLex.stopper}. We can write | |
| 49 | 451 | |
| 452 | @{ML_response [display]
 | |
| 453 | "let | |
| 50 | 454 | val input = filtered_input \"in | in | in\" | 
| 49 | 455 | in | 
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 456 | Scan.finite OuterLex.stopper | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 457 | (OuterParse.enum \"|\" (OuterParse.$$$ \"in\")) input | 
| 49 | 458 | end" | 
| 459 | "([\"in\",\"in\",\"in\"],[])"} | |
| 460 | ||
| 54 
1783211b3494
tuned; added document antiquotation ML_response_fake_both
 Christian Urban <urbanc@in.tum.de> parents: 
53diff
changeset | 461 | The following function will help us later to run examples | 
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 462 | |
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 463 | *} | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 464 | |
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 465 | ML {*
 | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 466 | fun parse p input = Scan.finite OuterLex.stopper (Scan.error p) input | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 467 | *} | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 468 | |
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 469 | text {*
 | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 470 | |
| 49 | 471 |   The function @{ML "OuterParse.!!!"} can be used to force termination of the
 | 
| 472 |   parser in case of a dead end, just like @{ML "Scan.!!"} (see previous section), 
 | |
| 473 |   except that the error message is fixed to be @{text [quotes] "Outer syntax error"}
 | |
| 474 | with a relatively precise description of the failure. For example: | |
| 475 | ||
| 476 | @{ML_response_fake [display]
 | |
| 477 | "let | |
| 50 | 478 | val input = filtered_input \"in |\" | 
| 49 | 479 | val parse_bar_then_in = OuterParse.$$$ \"|\" -- OuterParse.$$$ \"in\" | 
| 480 | in | |
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 481 | parse (OuterParse.!!! parse_bar_then_in) input | 
| 49 | 482 | end" | 
| 483 | "Exception ERROR \"Outer syntax error: keyword \"|\" expected, | |
| 484 | but keyword in was found\" raised" | |
| 485 | } | |
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 486 | |
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 487 |   \begin{exercise}
 | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 488 |   A type-identifier, for example @{typ "'a"}, is a token of 
 | 
| 54 
1783211b3494
tuned; added document antiquotation ML_response_fake_both
 Christian Urban <urbanc@in.tum.de> parents: 
53diff
changeset | 489 |   kind @{ML "Keyword" in OuterLex}. It can be parsed using 
 | 
| 
1783211b3494
tuned; added document antiquotation ML_response_fake_both
 Christian Urban <urbanc@in.tum.de> parents: 
53diff
changeset | 490 |   the function @{ML OuterParse.type_ident}.
 | 
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 491 |   \end{exercise}
 | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 492 | |
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 493 | |
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 494 | *} | 
| 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 495 | |
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 496 | |
| 49 | 497 | section {* Positional Information *}
 | 
| 498 | ||
| 499 | text {*
 | |
| 500 | ||
| 501 |   @{ML OuterParse.position}
 | |
| 502 | ||
| 503 | *} | |
| 504 | ||
| 505 | ML {*
 | |
| 506 | OuterParse.position | |
| 507 | *} | |
| 508 | ||
| 44 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 509 | |
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 510 | section {* Parsing Inner Syntax *}
 | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 511 | |
| 44 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 512 | ML {*
 | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 513 | let | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 514 | val input = OuterSyntax.scan Position.none "0" | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 515 | in | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 516 | OuterParse.prop input | 
| 
dee4b3e66dfe
added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 517 | end | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 518 | |
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 519 | *} | 
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 520 | |
| 50 | 521 | ML {*
 | 
| 522 | OuterParse.opt_target | |
| 523 | *} | |
| 524 | ||
| 525 | ML {*
 | |
| 526 | OuterParse.opt_target -- | |
| 527 | OuterParse.fixes -- | |
| 528 | OuterParse.for_fixes -- | |
| 529 | Scan.optional (OuterParse.$$$ "where" |-- | |
| 530 | OuterParse.!!! (OuterParse.enum1 "|" (SpecParse.opt_thm_name ":" -- OuterParse.prop))) [] | |
| 531 | ||
| 532 | *} | |
| 533 | ||
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 534 | text {* (FIXME funny output for a proposition) *}
 | 
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 535 | |
| 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 536 | |
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 537 | |
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 538 | chapter {* Parsing *}
 | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 539 | |
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 540 | text {*
 | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 541 | |
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 542 | Lots of Standard ML code is given in this document, for various reasons, | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 543 | including: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 544 |   \begin{itemize}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 545 | \item direct quotation of code found in the Isabelle source files, | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 546 | or simplified versions of such code | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 547 | \item identifiers found in the Isabelle source code, with their types | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 548 | (or specialisations of their types) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 549 | \item code examples, which can be run by the reader, to help illustrate the | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 550 | behaviour of functions found in the Isabelle source code | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 551 | \item ancillary functions, not from the Isabelle source code, | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 552 | which enable the reader to run relevant code examples | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 553 | \item type abbreviations, which help explain the uses of certain functions | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 554 |   \end{itemize}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 555 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 556 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 557 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 558 | section {* Parsing Isar input *}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 559 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 560 | text {*
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 561 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 562 | The typical parsing function has the type | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 563 |   \texttt{'src -> 'res * 'src}, with input  
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 564 |   of type \texttt{'src}, returning a result 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 565 |   of type \texttt{'res}, which is (or is derived from) the first part of the
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 566 | input, and also returning the remainder of the input. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 567 | (In the common case, when it is clear what the ``remainder of the input'' | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 568 | means, we will just say that the functions ``returns'' the | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 569 |   value of type \texttt{'res}). 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 570 | An exception is raised if an appropriate value | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 571 | cannot be produced from the input. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 572 | A range of exceptions can be used to identify different reasons | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 573 | for the failure of a parse. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 574 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 575 | This contrasts the standard parsing function in Standard ML, | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 576 | which is of type | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 577 |   \texttt{type ('res, 'src) reader = 'src -> ('res * 'src) option};
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 578 |   (for example, \texttt{List.getItem} and \texttt{Substring.getc}).
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 579 | However, much of the discussion at | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 580 | FIX file:/home/jeremy/html/ml/SMLBasis/string-cvt.html | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 581 | is relevant. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 582 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 583 | Naturally one may convert between the two different sorts of parsing functions | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 584 | as follows: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 585 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 586 | open StringCvt ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 587 |   type ('res, 'src) ex_reader = 'src -> 'res * 'src
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 588 |   (* ex_reader : ('res, 'src) reader -> ('res, 'src) ex_reader *)
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 589 | fun ex_reader rdr src = Option.valOf (rdr src) ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 590 |   (* reader : ('res, 'src) ex_reader -> ('res, 'src) reader *)
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 591 | fun reader exrdr src = SOME (exrdr src) handle _ => NONE ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 592 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 593 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 594 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 595 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 596 | section{* The \texttt{Scan} structure *}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 597 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 598 | text {* 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 599 |   The source file is \texttt{src/General/scan.ML}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 600 | This structure provides functions for using and combining parsing functions | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 601 |   of the type \texttt{'src -> 'res * 'src}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 602 | Three exceptions are used: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 603 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 604 | exception MORE of string option; (*need more input (prompt)*) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 605 | exception FAIL of string option; (*try alternatives (reason of failure)*) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 606 | exception ABORT of string; (*dead end*) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 607 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 608 | Many functions in this structure (generally those with names composed of | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 609 | symbols) are declared as infix. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 610 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 611 | Some functions from that structure are | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 612 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 613 |   |-- : ('src -> 'res1 * 'src') * ('src' -> 'res2 * 'src'') ->
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 614 | 'src -> 'res2 * 'src'' | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 615 |   --| : ('src -> 'res1 * 'src') * ('src' -> 'res2 * 'src'') ->
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 616 | 'src -> 'res1 * 'src'' | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 617 |   -- : ('src -> 'res1 * 'src') * ('src' -> 'res2 * 'src'') ->
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 618 |   'src -> ('res1 * 'res2) * 'src''
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 619 |   ^^ : ('src -> string * 'src') * ('src' -> string * 'src'') ->
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 620 | 'src -> string * 'src'' | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 621 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 622 | These functions parse a result off the input source twice. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 623 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 624 |   \texttt{|--} and \texttt{--|} 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 625 | return the first result and the second result, respectively. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 626 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 627 |   \texttt{--} returns both.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 628 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 629 | \verb|^^| returns the result of concatenating the two results | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 630 | (which must be strings). | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 631 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 632 | Note how, although the types | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 633 |   \texttt{'src}, \texttt{'src'} and \texttt{'src''} will normally be the same,
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 634 | the types as shown help suggest the behaviour of the functions. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 635 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 636 |   :-- : ('src -> 'res1 * 'src') * ('res1 -> 'src' -> 'res2 * 'src'') ->
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 637 |   'src -> ('res1 * 'res2) * 'src''
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 638 |   :|-- : ('src -> 'res1 * 'src') * ('res1 -> 'src' -> 'res2 * 'src'') ->
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 639 | 'src -> 'res2 * 'src'' | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 640 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 641 |   These are similar to \texttt{|--} and \texttt{--|},
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 642 | except that the second parsing function can depend on the result of the first. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 643 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 644 |   >> : ('src -> 'res1 * 'src') * ('res1 -> 'res2) -> 'src -> 'res2 * 'src'
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 645 |   || : ('src -> 'res_src) * ('src -> 'res_src) -> 'src -> 'res_src
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 646 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 647 |   \texttt{p >> f} applies a function \texttt{f} to the result of a parse.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 648 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 649 |   \texttt{||} tries a second parsing function if the first one
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 650 |   fails by raising an exception of the form \texttt{FAIL \_}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 651 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 652 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 653 |   succeed : 'res -> ('src -> 'res * 'src) ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 654 |   fail : ('src -> 'res_src) ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 655 |   !! : ('src * string option -> string) -> 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 656 |   ('src -> 'res_src) -> ('src -> 'res_src) ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 657 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 658 |   \texttt{succeed r} returns \texttt{r}, with the input unchanged.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 659 |   \texttt{fail} always fails, raising exception \texttt{FAIL NONE}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 660 |   \texttt{!! f} only affects the failure mode, turning a failure that
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 661 |   raises \texttt{FAIL \_} into a failure that raises \texttt{ABORT ...}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 662 | This is used to prevent recovery from the failure --- | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 663 |   thus, in \texttt{!! parse1 || parse2}, if \texttt{parse1} fails, 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 664 |   it won't recover by trying \texttt{parse2}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 665 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 666 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 667 |   one : ('si -> bool) -> ('si list -> 'si * 'si list) ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 668 |   some : ('si -> 'res option) -> ('si list -> 'res * 'si list) ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 669 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 670 | These require the input to be a list of items: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 671 |   they fail, raising \texttt{MORE NONE} if the list is empty.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 672 |   On other failures they raise \texttt{FAIL NONE} 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 673 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 674 |   \texttt{one p} takes the first
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 675 |   item from the list if it satisfies \texttt{p}, otherwise fails.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 676 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 677 |   \texttt{some f} takes the first
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 678 |   item from the list and applies \texttt{f} to it, failing if this returns
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 679 |   \texttt{NONE}.  
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 680 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 681 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 682 |   many : ('si -> bool) -> 'si list -> 'si list * 'si list ; 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 683 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 684 |   \texttt{many p} takes items from the input until it encounters one 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 685 |   which does not satisfy \texttt{p}.  If it reaches the end of the input
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 686 |   it fails, raising \texttt{MORE NONE}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 687 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 688 |   \texttt{many1} (with the same type) fails if the first item 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 689 |   does not satisfy \texttt{p}.  
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 690 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 691 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 692 |   option : ('src -> 'res * 'src) -> ('src -> 'res option * 'src)
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 693 |   optional : ('src -> 'res * 'src) -> 'res -> ('src -> 'res * 'src)
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 694 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 695 |   \texttt{option}: 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 696 |   where the parser \texttt{f} succeeds with result \texttt{r} 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 697 |   or raises \texttt{FAIL \_},
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 698 |   \texttt{option f} gives the result \texttt{SOME r} or \texttt{NONE}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 699 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 700 |   \texttt{optional}: if parser \texttt{f} fails by raising \texttt{FAIL \_},
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 701 |   \texttt{optional f default} provides the result \texttt{default}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 702 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 703 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 704 |   repeat : ('src -> 'res * 'src) -> 'src -> 'res list * 'src
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 705 |   repeat1 : ('src -> 'res * 'src) -> 'src -> 'res list * 'src
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 706 |   bulk : ('src -> 'res * 'src) -> 'src -> 'res list * 'src 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 707 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 708 |   \texttt{repeat f} repeatedly parses an item off the remaining input until 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 709 |   \texttt{f} fails with \texttt{FAIL \_}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 710 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 711 |   \texttt{repeat1} is as for \texttt{repeat}, but requires at least one
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 712 | successful parse. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 713 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 714 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 715 |   lift : ('src -> 'res * 'src) -> ('ex * 'src -> 'res * ('ex * 'src))
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 716 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 717 |   \texttt{lift} changes the source type of a parser by putting in an extra
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 718 |   component \texttt{'ex}, which is ignored in the parsing.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 719 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 720 |   The \texttt{Scan} structure also provides the type \texttt{lexicon}, 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 721 | HOW DO THEY WORK ?? TO BE COMPLETED | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 722 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 723 | dest_lexicon: lexicon -> string list ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 724 | make_lexicon: string list list -> lexicon ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 725 | empty_lexicon: lexicon ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 726 | extend_lexicon: string list list -> lexicon -> lexicon ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 727 | merge_lexicons: lexicon -> lexicon -> lexicon ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 728 | is_literal: lexicon -> string list -> bool ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 729 | literal: lexicon -> string list -> string list * string list ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 730 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 731 | Two lexicons, for the commands and keywords, are stored and can be retrieved | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 732 | by: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 733 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 734 | val (command_lexicon, keyword_lexicon) = OuterSyntax.get_lexicons () ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 735 | val commands = Scan.dest_lexicon command_lexicon ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 736 | val keywords = Scan.dest_lexicon keyword_lexicon ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 737 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 738 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 739 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 740 | section{* The \texttt{OuterLex} structure *}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 741 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 742 | text {*
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 743 |   The source file is @{text "src/Pure/Isar/outer_lex.ML"}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 744 | In some other source files its name is abbreviated: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 745 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 746 | structure T = OuterLex; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 747 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 748 |   This structure defines the type \texttt{token}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 749 | (The types | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 750 |   \texttt{OuterLex.token},
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 751 |   \texttt{OuterParse.token} and
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 752 |   \texttt{SpecParse.token} are all the same).
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 753 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 754 | Input text is split up into tokens, and the input source type for many parsing | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 755 |   functions is \texttt{token list}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 756 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 757 | The datatype definition (which is not published in the signature) is | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 758 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 759 | datatype token = Token of Position.T * (token_kind * string); | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 760 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 761 | but here are some runnable examples for viewing tokens: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 762 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 763 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 764 | |
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 765 | |
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 766 | |
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 767 | |
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 768 | ML {*
 | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 769 | val toks = OuterSyntax.scan Position.none | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 770 |    "theory,imports;begin x.y.z apply ?v1 ?'a 'a -- || 44 simp (* xx *) { * fff * }" ;
 | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 771 | *} | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 772 | |
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 773 | ML {*
 | 
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 774 | print_depth 20 ; | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 775 | *} | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 776 | |
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 777 | ML {*
 | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 778 | map OuterLex.text_of toks ; | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 779 | *} | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 780 | |
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 781 | ML {*
 | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 782 | val proper_toks = filter OuterLex.is_proper toks ; | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 783 | *} | 
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 784 | |
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 785 | ML {*
 | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 786 | map OuterLex.kind_of proper_toks | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 787 | *} | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 788 | |
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 789 | ML {*
 | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 790 | map OuterLex.unparse proper_toks ; | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 791 | *} | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 792 | |
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 793 | ML {*
 | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 794 | OuterLex.stopper | 
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 795 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 796 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 797 | text {*
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 798 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 799 |   The function \texttt{is\_proper : token -> bool} identifies tokens which are
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 800 | not white space or comments: many parsing functions assume require spaces or | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 801 | comments to have been filtered out. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 802 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 803 | There is a special end-of-file token: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 804 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 805 | val (tok_eof : token, is_eof : token -> bool) = T.stopper ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 806 | (* end of file token *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 807 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 808 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 809 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 810 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 811 | section {* The \texttt{OuterParse} structure *}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 812 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 813 | text {*
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 814 |   The source file is \texttt{src/Pure/Isar/outer\_parse.ML}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 815 | In some other source files its name is abbreviated: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 816 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 817 | structure P = OuterParse; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 818 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 819 |   Here the parsers use \texttt{token list} as the input source type. 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 820 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 821 | Some of the parsers simply select the first token, provided that it is of the | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 822 |   right kind (as returned by \texttt{T.kind\_of}): these are 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 823 |   \texttt{ command, keyword, short\_ident, long\_ident, sym\_ident, term\_var,
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 824 | type\_ident, type\_var, number, string, alt\_string, verbatim, sync, eof} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 825 | Others select the first token, provided that it is one of several kinds, | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 826 |   (eg, \texttt{name, xname, text, typ}).
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 827 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 828 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 829 | type 'a tlp = token list -> 'a * token list ; (* token list parser *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 830 | $$$ : string -> string tlp | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 831 | nat : int tlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 832 | maybe : 'a tlp -> 'a option tlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 833 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 834 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 835 |   \texttt{\$\$\$ s} returns the first token,
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 836 |   if it equals \texttt{s} \emph{and} \texttt{s} is a keyword.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 837 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 838 |   \texttt{nat} returns the first token, if it is a number, and evaluates it.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 839 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 840 |   \texttt{maybe}: if \texttt{p} returns \texttt{r}, 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 841 |   then \texttt{maybe p} returns \texttt{SOME r} ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 842 |   if the first token is an underscore, it returns \texttt{NONE}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 843 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 844 | A few examples: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 845 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 846 | P.list : 'a tlp -> 'a list tlp ; (* likewise P.list1 *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 847 | P.and_list : 'a tlp -> 'a list tlp ; (* likewise P.and_list1 *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 848 | val toks : token list = OuterSyntax.scan "44 ,_, 66,77" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 849 | val proper_toks = List.filter T.is_proper toks ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 850 | P.list P.nat toks ; (* OK, doesn't recognize white space *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 851 | P.list P.nat proper_toks ; (* fails, doesn't recognize what follows ',' *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 852 | P.list (P.maybe P.nat) proper_toks ; (* fails, end of input *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 853 | P.list (P.maybe P.nat) (proper_toks @ [tok_eof]) ; (* OK *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 854 | val toks : token list = OuterSyntax.scan "44 and 55 and 66 and 77" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 855 | P.and_list P.nat (List.filter T.is_proper toks @ [tok_eof]) ; (* ??? *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 856 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 857 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 858 | The following code helps run examples: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 859 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 860 | fun parse_str tlp str = | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 861 | let val toks : token list = OuterSyntax.scan str ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 862 | val proper_toks = List.filter T.is_proper toks @ [tok_eof] ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 863 | val (res, rem_toks) = tlp proper_toks ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 864 | val rem_str = String.concat | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 865 | (Library.separate " " (List.map T.unparse rem_toks)) ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 866 | in (res, rem_str) end ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 867 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 868 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 869 |   Some examples from \texttt{src/Pure/Isar/outer\_parse.ML}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 870 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 871 | val type_args = | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 872 | type_ident >> Library.single || | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 873 |   $$$ "(" |-- !!! (list1 type_ident --| $$$ ")") ||
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 874 | Scan.succeed []; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 875 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 876 | There are three ways parsing a list of type arguments can succeed. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 877 | The first line reads a single type argument, and turns it into a singleton | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 878 | list. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 879 |   The second line reads "(", and then the remainder, ignoring the "(" ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 880 | the remainder consists of a list of type identifiers (at least one), | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 881 | and then a ")" which is also ignored. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 882 |   The \texttt{!!!} ensures that if the parsing proceeds this far and then fails,
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 883 |   it won't try the third line (see the description of \texttt{Scan.!!}).
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 884 | The third line consumes no input and returns the empty list. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 885 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 886 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 887 | fun triple2 (x, (y, z)) = (x, y, z); | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 888 | val arity = xname -- ($$$ "::" |-- !!! ( | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 889 |   Scan.optional ($$$ "(" |-- !!! (list1 sort --| $$$ ")")) []
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 890 | -- sort)) >> triple2; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 891 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 892 |   The parser \texttt{arity} reads a typename $t$, then ``\texttt{::}'' (which is
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 893 | ignored), then optionally a list $ss$ of sorts and then another sort $s$. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 894 |   The result $(t, (ss, s))$ is transformed by \texttt{triple2} to $(t, ss, s)$.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 895 | The second line reads the optional list of sorts: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 896 |   it reads first ``\texttt{(}'' and last ``\texttt{)}'', which are both ignored,
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 897 | and between them a comma-separated list of sorts. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 898 |   If this list is absent, the default \texttt{[]} provides the list of sorts.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 899 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 900 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 901 |   parse_str P.type_args "('a, 'b) ntyp" ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 902 | parse_str P.type_args "'a ntyp" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 903 | parse_str P.type_args "ntyp" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 904 | parse_str P.arity "ty :: tycl" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 905 | parse_str P.arity "ty :: (tycl1, tycl2) tycl" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 906 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 907 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 908 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 909 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 910 | section {* The \texttt{SpecParse} structure *}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 911 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 912 | text {*
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 913 |   The source file is \texttt{src/Pure/Isar/spec\_parse.ML}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 914 | This structure contains token list parsers for more complicated values. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 915 | For example, | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 916 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 917 | open SpecParse ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 918 | attrib : Attrib.src tok_rdr ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 919 | attribs : Attrib.src list tok_rdr ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 920 | opt_attribs : Attrib.src list tok_rdr ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 921 | xthm : (thmref * Attrib.src list) tok_rdr ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 922 | xthms1 : (thmref * Attrib.src list) list tok_rdr ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 923 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 924 | parse_str attrib "simp" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 925 | parse_str opt_attribs "hello" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 926 | val (ass, "") = parse_str attribs "[standard, xxxx, simp, intro, OF sym]" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 927 | map Args.dest_src ass ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 928 | val (asrc, "") = parse_str attrib "THEN trans [THEN sym]" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 929 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 930 | parse_str xthm "mythm [attr]" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 931 | parse_str xthms1 "thm1 [attr] thms2" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 932 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 933 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 934 |   As you can see, attributes are described using types of the \texttt{Args}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 935 | structure, described below. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 936 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 937 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 938 | section{* The \texttt{Args} structure *}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 939 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 940 | text {*
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 941 |   The source file is \texttt{src/Pure/Isar/args.ML}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 942 |   The primary type of this structure is the \texttt{src} datatype;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 943 | the single constructors not published in the signature, but | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 944 |   \texttt{Args.src} and \texttt{Args.dest\_src}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 945 | are in fact the constructor and destructor functions. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 946 |   Note that the types \texttt{Attrib.src} and \texttt{Method.src}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 947 |   are in fact \texttt{Args.src}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 948 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 949 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 950 | src : (string * Args.T list) * Position.T -> Args.src ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 951 | dest_src : Args.src -> (string * Args.T list) * Position.T ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 952 | Args.pretty_src : Proof.context -> Args.src -> Pretty.T ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 953 | fun pr_src ctxt src = Pretty.string_of (Args.pretty_src ctxt src) ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 954 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 955 | val thy = ML_Context.the_context () ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 956 | val ctxt = ProofContext.init thy ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 957 | map (pr_src ctxt) ass ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 958 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 959 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 960 |   So an \texttt{Args.src} consists of the first word, then a list of further 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 961 |   ``arguments'', of type \texttt{Args.T}, with information about position in the
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 962 | input. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 963 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 964 | (* how an Args.src is parsed *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 965 |   P.position : 'a tlp -> ('a * Position.T) tlp ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 966 | P.arguments : Args.T list tlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 967 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 968 | val parse_src : Args.src tlp = | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 969 | P.position (P.xname -- P.arguments) >> Args.src ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 970 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 971 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 972 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 973 | val ((first_word, args), pos) = Args.dest_src asrc ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 974 | map Args.string_of args ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 975 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 976 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 977 |   The \texttt{Args} structure contains more parsers and parser transformers 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 978 |   for which the input source type is \texttt{Args.T list}.  For example,
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 979 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 980 | type 'a atlp = Args.T list -> 'a * Args.T list ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 981 | open Args ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 982 | nat : int atlp ; (* also Args.int *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 983 | thm_sel : PureThy.interval list atlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 984 | list : 'a atlp -> 'a list atlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 985 | attribs : (string -> string) -> Args.src list atlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 986 | opt_attribs : (string -> string) -> Args.src list atlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 987 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 988 | (* parse_atl_str : 'a atlp -> (string -> 'a * string) ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 989 | given an Args.T list parser, to get a string parser *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 990 | fun parse_atl_str atlp str = | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 991 | let val (ats, rem_str) = parse_str P.arguments str ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 992 | val (res, rem_ats) = atlp ats ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 993 | in (res, String.concat (Library.separate " " | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 994 | (List.map Args.string_of rem_ats @ [rem_str]))) end ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 995 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 996 | parse_atl_str Args.int "-1-," ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 997 | parse_atl_str (Scan.option Args.int) "x1-," ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 998 | parse_atl_str Args.thm_sel "(1-,4,13-22)" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 999 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1000 | val (ats as atsrc :: _, "") = parse_atl_str (Args.attribs I) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1001 | "[THEN trans [THEN sym], simp, OF sym]" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1002 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1003 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1004 |   From here, an attribute is interpreted using \texttt{Attrib.attribute}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1005 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1006 |   \texttt{Args} has a large number of functions which parse an \texttt{Args.src}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1007 | and also refer to a generic context. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1008 |   Note the use of \texttt{Scan.lift} for this.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1009 |   (as does \texttt{Attrib} - RETHINK THIS)
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1010 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1011 |   (\texttt{Args.syntax} shown below has type specialised)
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1012 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1013 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1014 |   type ('res, 'src) parse_fn = 'src -> 'res * 'src ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1015 |   type 'a cgatlp = ('a, Context.generic * Args.T list) parse_fn ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1016 | Scan.lift : 'a atlp -> 'a cgatlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1017 | term : term cgatlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1018 | typ : typ cgatlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1019 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1020 |   Args.syntax : string -> 'res cgatlp -> src -> ('res, Context.generic) parse_fn ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1021 | Attrib.thm : thm cgatlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1022 | Attrib.thms : thm list cgatlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1023 | Attrib.multi_thm : thm list cgatlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1024 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1025 | (* parse_cgatl_str : 'a cgatlp -> (string -> 'a * string) ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1026 | given a (Context.generic * Args.T list) parser, to get a string parser *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1027 | fun parse_cgatl_str cgatlp str = | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1028 | let | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1029 | (* use the current generic context *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1030 | val generic = Context.Theory thy ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1031 | val (ats, rem_str) = parse_str P.arguments str ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1032 | (* ignore any change to the generic context *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1033 | val (res, (_, rem_ats)) = cgatlp (generic, ats) ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1034 | in (res, String.concat (Library.separate " " | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1035 | (List.map Args.string_of rem_ats @ [rem_str]))) end ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1036 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1037 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1038 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1039 | section{* Attributes, and the \texttt{Attrib} structure *}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1040 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1041 | text {*
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1042 |   The type \texttt{attribute} is declared in \texttt{src/Pure/thm.ML}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1043 |   The source file for the \texttt{Attrib} structure is
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1044 |   \texttt{src/Pure/Isar/attrib.ML}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1045 | Most attributes use a theorem to change a generic context (for example, | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1046 | by declaring that the theorem should be used, by default, in simplification), | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1047 | or change a theorem (which most often involves referring to the current | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1048 | theory). | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1049 |   The functions \texttt{Thm.rule\_attribute} and
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1050 |   \texttt{Thm.declaration\_attribute} create attributes of these kinds.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1051 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1052 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1053 | type attribute = Context.generic * thm -> Context.generic * thm; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1054 | type 'a trf = 'a -> 'a ; (* transformer of a given type *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1055 | Thm.rule_attribute : (Context.generic -> thm -> thm) -> attribute ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1056 | Thm.declaration_attribute : (thm -> Context.generic trf) -> attribute ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1057 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1058 | Attrib.print_attributes : theory -> unit ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1059 | Attrib.pretty_attribs : Proof.context -> src list -> Pretty.T list ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1060 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1061 | List.app Pretty.writeln (Attrib.pretty_attribs ctxt ass) ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1062 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1063 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1064 | An attribute is stored in a theory as indicated by: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1065 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1066 | Attrib.add_attributes : | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1067 | (bstring * (src -> attribute) * string) list -> theory trf ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1068 | (* | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1069 |   Attrib.add_attributes [("THEN", THEN_att, "resolution with rule")] ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1070 | *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1071 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1072 | where the first and third arguments are name and description of the attribute, | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1073 | and the second is a function which parses the attribute input text | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1074 | (including the attribute name, which has necessarily already been parsed). | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1075 |   Here, \texttt{THEN\_att} is a function declared in the code for the
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1076 |   structure \texttt{Attrib}, but not published in its signature.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1077 |   The source file \texttt{src/Pure/Isar/attrib.ML} shows the use of 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1078 |   \texttt{Attrib.add\_attributes} to add a number of attributes.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1079 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1080 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1081 | FullAttrib.THEN_att : src -> attribute ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1082 | FullAttrib.THEN_att atsrc (generic, ML_Context.thm "sym") ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1083 | FullAttrib.THEN_att atsrc (generic, ML_Context.thm "all_comm") ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1084 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1085 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1086 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1087 | Attrib.syntax : attribute cgatlp -> src -> attribute ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1088 | Attrib.no_args : attribute -> src -> attribute ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1089 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1090 |   When this is called as \texttt{syntax scan src (gc, th)}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1091 |   the generic context \texttt{gc} is used 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1092 |   (and potentially changed to \texttt{gc'})
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1093 |   by \texttt{scan} in parsing to obtain an attribute \texttt{attr} which would
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1094 |   then be applied to \texttt{(gc', th)}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1095 |   The source for parsing the attribute is the arguments part of \texttt{src},
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1096 | which must all be consumed by the parse. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1097 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1098 |   For example, for \texttt{Attrib.no\_args attr src}, the attribute parser 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1099 |   simply returns \texttt{attr}, requiring that the arguments part of
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1100 |   \texttt{src} must be empty.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1101 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1102 |   Some examples from \texttt{src/Pure/Isar/attrib.ML}, modified:
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1103 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1104 | fun rot_att_n n (gc, th) = (gc, rotate_prems n th) ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1105 | rot_att_n : int -> attribute ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1106 | val rot_arg = Scan.lift (Scan.optional Args.int 1 : int atlp) : int cgatlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1107 | val rotated_att : src -> attribute = | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1108 | Attrib.syntax (rot_arg >> rot_att_n : attribute cgatlp) ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1109 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1110 | val THEN_arg : int cgatlp = Scan.lift | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1111 | (Scan.optional (Args.bracks Args.nat : int atlp) 1 : int atlp) ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1112 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1113 | Attrib.thm : thm cgatlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1114 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1115 | THEN_arg -- Attrib.thm : (int * thm) cgatlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1116 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1117 | fun THEN_att_n (n, tht) (gc, th) = (gc, th RSN (n, tht)) ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1118 | THEN_att_n : int * thm -> attribute ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1119 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1120 | val THEN_att : src -> attribute = Attrib.syntax | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1121 | (THEN_arg -- Attrib.thm >> THEN_att_n : attribute cgatlp); | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1122 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1123 |   The functions I've called \texttt{rot\_arg} and \texttt{THEN\_arg}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1124 |   read an optional argument, which for \texttt{rotated} is an integer, 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1125 |   and for \texttt{THEN} is a natural enclosed in square brackets;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1126 | the default, if the argument is absent, is 1 in each case. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1127 |   Functions \texttt{rot\_att\_n} and \texttt{THEN\_att\_n} turn these into
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1128 |   attributes, where \texttt{THEN\_att\_n} also requires a theorem, which is
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1129 |   parsed by \texttt{Attrib.thm}.  
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1130 |   Infix operators \texttt{--} and \texttt{>>} are in the structure \texttt{Scan}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1131 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1132 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1133 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1134 | section{* Methods, and the \texttt{Method} structure *}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1135 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1136 | text {*
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1137 |   The source file is \texttt{src/Pure/Isar/method.ML}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1138 |   The type \texttt{method} is defined by the datatype declaration
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1139 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1140 | (* datatype method = Meth of thm list -> cases_tactic; *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1141 | RuleCases.NO_CASES : tactic -> cases_tactic ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1142 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1143 |   In fact \texttt{RAW\_METHOD\_CASES} (below) is exactly the constructor
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1144 |   \texttt{Meth}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1145 |   A \texttt{cases\_tactic} is an elaborated version of a tactic.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1146 |   \texttt{NO\_CASES tac} is a \texttt{cases\_tactic} which consists of a
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1147 |   \texttt{cases\_tactic} without any further case information.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1148 |   For further details see the description of structure \texttt{RuleCases} below.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1149 | The list of theorems to be passed to a method consists of the current | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1150 |   \emph{facts} in the proof.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1151 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1152 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1153 | RAW_METHOD : (thm list -> tactic) -> method ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1154 | METHOD : (thm list -> tactic) -> method ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1155 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1156 | SIMPLE_METHOD : tactic -> method ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1157 | SIMPLE_METHOD' : (int -> tactic) -> method ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1158 | SIMPLE_METHOD'' : ((int -> tactic) -> tactic) -> (int -> tactic) -> method ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1159 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1160 | RAW_METHOD_CASES : (thm list -> cases_tactic) -> method ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1161 | METHOD_CASES : (thm list -> cases_tactic) -> method ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1162 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1163 | A method is, in its simplest form, a tactic; applying the method is to apply | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1164 | the tactic to the current goal state. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1165 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1166 |   Applying \texttt{RAW\_METHOD tacf} creates a tactic by applying 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1167 |   \texttt{tacf} to the current {facts}, and applying that tactic to the
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1168 | goal state. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1169 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1170 |   \texttt{METHOD} is similar but also first applies
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1171 |   \texttt{Goal.conjunction\_tac} to all subgoals.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1172 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1173 |   \texttt{SIMPLE\_METHOD tac} inserts the facts into all subgoals and then
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1174 |   applies \texttt{tacf}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1175 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1176 |   \texttt{SIMPLE\_METHOD' tacf} inserts the facts and then
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1177 |   applies \texttt{tacf} to subgoal 1.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1178 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1179 |   \texttt{SIMPLE\_METHOD'' quant tacf} does this for subgoal(s) selected by
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1180 |   \texttt{quant}, which may be, for example,
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1181 |   \texttt{ALLGOALS} (all subgoals),
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1182 |   \texttt{TRYALL} (try all subgoals, failure is OK),
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1183 |   \texttt{FIRSTGOAL} (try subgoals until it succeeds once), 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1184 |   \texttt{(fn tacf => tacf 4)} (subgoal 4), etc
 | 
| 16 | 1185 |   (see the \texttt{Tactical} structure, FIXME) %%\cite[Chapter 4]{ref}).
 | 
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1186 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1187 | A method is stored in a theory as indicated by: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1188 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1189 | Method.add_method : | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1190 | (bstring * (src -> Proof.context -> method) * string) -> theory trf ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1191 | ( * | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1192 | * ) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1193 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1194 | where the first and third arguments are name and description of the method, | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1195 | and the second is a function which parses the method input text | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1196 | (including the method name, which has necessarily already been parsed). | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1197 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1198 |   Here, \texttt{xxx} is a function declared in the code for the
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1199 |   structure \texttt{Method}, but not published in its signature.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1200 |   The source file \texttt{src/Pure/Isar/method.ML} shows the use of 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1201 |   \texttt{Method.add\_method} to add a number of methods.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1202 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1203 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1204 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1205 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1206 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1207 | end |