| author | Christian Urban <urbanc@in.tum.de> | 
| Tue, 27 Oct 2009 15:43:21 +0100 | |
| changeset 361 | 8ba963a3e039 | 
| parent 357 | 80b56d9b322f | 
| child 366 | 06f044466f24 | 
| permissions | -rw-r--r-- | 
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1 | theory Parsing | 
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 2 | imports Base "Helper/Command/Command" "Package/Simple_Inductive_Package" | 
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 3 | begin | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 4 | |
| 346 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
344diff
changeset | 5 | (*<*) | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
344diff
changeset | 6 | setup {*
 | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
344diff
changeset | 7 | open_file_with_prelude | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
344diff
changeset | 8 | "Parsing_Code.thy" | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
344diff
changeset | 9 | ["theory Parsing", | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
344diff
changeset | 10 | "imports Base \"Package/Simple_Inductive_Package\"", | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
344diff
changeset | 11 | "begin"] | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
344diff
changeset | 12 | *} | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
344diff
changeset | 13 | (*>*) | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
344diff
changeset | 14 | |
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 15 | chapter {* Parsing *}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 16 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 17 | text {*
 | 
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 18 |   Isabelle distinguishes between \emph{outer} and \emph{inner}
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 19 |   syntax. Commands, such as \isacommand{definition}, \isacommand{inductive}
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 20 | and so on, belong to the outer syntax, whereas terms, types and so on belong | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 21 | to the inner syntax. For parsing inner syntax, Isabelle uses a rather | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 22 | general and sophisticated algorithm, which is driven by priority | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 23 | grammars. Parsers for outer syntax are built up by functional parsing | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 24 | combinators. These combinators are a well-established technique for parsing, | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 25 | which has, for example, been described in Paulson's classic ML-book | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 26 |   \cite{paulson-ml2}.  Isabelle developers are usually concerned with writing
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 27 | these outer syntax parsers, either for new definitional packages or for | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 28 | calling methods with specific arguments. | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 29 | |
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 30 |   \begin{readmore}
 | 
| 236 | 31 | The library for writing parser combinators is split up, roughly, into two | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 32 | parts: The first part consists of a collection of generic parser combinators | 
| 236 | 33 |   defined in the structure @{ML_struct Scan} in the file @{ML_file
 | 
| 34 | "Pure/General/scan.ML"}. The second part of the library consists of | |
| 35 | combinators for dealing with specific token types, which are defined in the | |
| 36 |   structure @{ML_struct OuterParse} in the file @{ML_file
 | |
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 37 | "Pure/Isar/outer_parse.ML"}. In addition specific parsers for packages are | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 38 |   defined in @{ML_file "Pure/Isar/spec_parse.ML"}. Parsers for method arguments 
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 39 |   are defined in @{ML_file "Pure/Isar/args.ML"}.
 | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 40 |   \end{readmore}
 | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 41 | |
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 42 | *} | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 43 | |
| 49 | 44 | section {* Building Generic Parsers *}
 | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 45 | |
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 46 | text {*
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 47 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 48 | Let us first have a look at parsing strings using generic parsing | 
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 49 |   combinators. The function @{ML_ind "$$" in Scan} takes a string as argument and will
 | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 50 | ``consume'' this string from a given input list of strings. ``Consume'' in | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 51 | this context means that it will return a pair consisting of this string and | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 52 | the rest of the input list. For example: | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 53 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 54 |   @{ML_response [display,gray] 
 | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 55 | "($$ \"h\") (Symbol.explode \"hello\")" "(\"h\", [\"e\", \"l\", \"l\", \"o\"])"} | 
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 56 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 57 |   @{ML_response [display,gray] 
 | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 58 | "($$ \"w\") (Symbol.explode \"world\")" "(\"w\", [\"o\", \"r\", \"l\", \"d\"])"} | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 59 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 60 |   The function @{ML "$$"} will either succeed (as in the two examples above)
 | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 61 |   or raise the exception @{text "FAIL"} if no string can be consumed. For
 | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 62 | example trying to parse | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 63 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 64 |   @{ML_response_fake [display,gray] 
 | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 65 | "($$ \"x\") (Symbol.explode \"world\")" | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 66 | "Exception FAIL raised"} | 
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 67 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 68 |   will raise the exception @{text "FAIL"}.  There are three exceptions used in
 | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 69 | the parsing combinators: | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 70 | |
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 71 |   \begin{itemize}
 | 
| 58 | 72 |   \item @{text "FAIL"} is used to indicate that alternative routes of parsing 
 | 
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 73 | might be explored. | 
| 58 | 74 |   \item @{text "MORE"} indicates that there is not enough input for the parser. For example 
 | 
| 75 |   in @{text "($$ \"h\") []"}.
 | |
| 60 
5b9c6010897b
doem tuning and made the cookbook work again with recent changes (CookBook/Package/Ind_Interface.thy needs to be looked at to see what the problem with the new parser type is)
 Christian Urban <urbanc@in.tum.de> parents: 
59diff
changeset | 76 |   \item @{text "ABORT"} is the exception that is raised when a dead end is reached. 
 | 
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 77 |   It is used for example in the function @{ML "!!"} (see below).
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 78 |   \end{itemize}
 | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 79 | |
| 50 | 80 | However, note that these exceptions are private to the parser and cannot be accessed | 
| 49 | 81 | by the programmer (for example to handle them). | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 82 | |
| 357 
80b56d9b322f
included a comment from Tim Bourke
 Christian Urban <urbanc@in.tum.de> parents: 
346diff
changeset | 83 |   In the examples above we use the function @{ML_ind explode in Symbol} from the
 | 
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 84 |   structure @{ML_struct Symbol}, instead of the more standard library function
 | 
| 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 85 |   @{ML_ind explode}, for obtaining an input list for the parser. The reason is
 | 
| 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 86 |   that @{ML_ind explode} in @{ML_struct Symbol} is aware of character
 | 
| 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 87 |   sequences, for example @{text "\<foo>"}, that have a special meaning in
 | 
| 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 88 | Isabelle. To see the difference consider | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 89 | |
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 90 | @{ML_response_fake [display,gray]
 | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 91 | "let | 
| 261 | 92 | val input = \"\<foo> bar\" | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 93 | in | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 94 | (explode input, Symbol.explode input) | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 95 | end" | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 96 | "([\"\\\", \"<\", \"f\", \"o\", \"o\", \">\", \" \", \"b\", \"a\", \"r\"], | 
| 261 | 97 | [\"\<foo>\", \" \", \"b\", \"a\", \"r\"])"} | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 98 | |
| 256 
1fb8d62c88a0
added some first index-information
 Christian Urban <urbanc@in.tum.de> parents: 
250diff
changeset | 99 |   Slightly more general than the parser @{ML "$$"} is the function 
 | 
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 100 |   @{ML_ind one in Scan}, in that it takes a predicate as argument and 
 | 
| 256 
1fb8d62c88a0
added some first index-information
 Christian Urban <urbanc@in.tum.de> parents: 
250diff
changeset | 101 | then parses exactly | 
| 52 | 102 | one item from the input list satisfying this predicate. For example the | 
| 58 | 103 |   following parser either consumes an @{text [quotes] "h"} or a @{text
 | 
| 49 | 104 | [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 | 105 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 106 | @{ML_response [display,gray] 
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 107 | "let | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 108 | val hw = Scan.one (fn x => x = \"h\" orelse x = \"w\") | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 109 | val input1 = Symbol.explode \"hello\" | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 110 | val input2 = Symbol.explode \"world\" | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 111 | in | 
| 236 | 112 | (hw input1, hw input2) | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 113 | end" | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 114 | "((\"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 | 115 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 116 |   Two parsers can be connected in sequence by using the function @{ML_ind "--" in Scan}. 
 | 
| 220 | 117 |   For example parsing @{text "h"}, @{text "e"} and @{text "l"} (in this 
 | 
| 118 | order) you can achieve by: | |
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 119 | |
| 236 | 120 |   @{ML_response [display,gray] 
 | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 121 | "($$ \"h\" -- $$ \"e\" -- $$ \"l\") (Symbol.explode \"hello\")" | 
| 236 | 122 | "(((\"h\", \"e\"), \"l\"), [\"l\", \"o\"])"} | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 123 | |
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 124 | 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 | 125 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 126 | If, as in the previous example, you want to parse a particular string, | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 127 |   then you can use the function @{ML_ind this_string in Scan}.
 | 
| 56 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 128 | |
| 236 | 129 |   @{ML_response [display,gray] 
 | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 130 | "Scan.this_string \"hell\" (Symbol.explode \"hello\")" | 
| 236 | 131 | "(\"hell\", [\"o\"])"} | 
| 56 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 132 | |
| 256 
1fb8d62c88a0
added some first index-information
 Christian Urban <urbanc@in.tum.de> parents: 
250diff
changeset | 133 | Parsers that explore alternatives can be constructed using the function | 
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 134 |   @{ML_ind  "||" in Scan}. The parser @{ML "(p || q)" for p q} returns the
 | 
| 58 | 135 |   result of @{text "p"}, in case it succeeds, otherwise it returns the
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 136 |   result of @{text "q"}. For example:
 | 
| 56 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 137 | |
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 138 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 139 | @{ML_response [display,gray] 
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 140 | "let | 
| 236 | 141 | val hw = $$ \"h\" || $$ \"w\" | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 142 | val input1 = Symbol.explode \"hello\" | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 143 | val input2 = Symbol.explode \"world\" | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 144 | in | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 145 | (hw input1, hw input2) | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 146 | end" | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 147 | "((\"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 | 148 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 149 |   The functions @{ML_ind "|--" in Scan} and @{ML_ind "--|" in Scan} work like the sequencing
 | 
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 150 | function for parsers, except that they discard the item being parsed by the | 
| 357 
80b56d9b322f
included a comment from Tim Bourke
 Christian Urban <urbanc@in.tum.de> parents: 
346diff
changeset | 151 | first (respectively second) parser. That means the item being dropped is the | 
| 
80b56d9b322f
included a comment from Tim Bourke
 Christian Urban <urbanc@in.tum.de> parents: 
346diff
changeset | 152 |   one that @{ML_ind "|--" in Scan} and @{ML_ind "--|" in Scan} ``point'' away.
 | 
| 
80b56d9b322f
included a comment from Tim Bourke
 Christian Urban <urbanc@in.tum.de> parents: 
346diff
changeset | 153 | For example: | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 154 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 155 | @{ML_response [display,gray]
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 156 | "let | 
| 236 | 157 | val just_e = $$ \"h\" |-- $$ \"e\" | 
| 158 | val just_h = $$ \"h\" --| $$ \"e\" | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 159 | val input = Symbol.explode \"hello\" | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 160 | in | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 161 | (just_e input, just_h input) | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 162 | end" | 
| 241 | 163 | "((\"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 | 164 | |
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 165 |   The parser @{ML "Scan.optional p x" for p x} returns the result of the parser 
 | 
| 58 | 166 |   @{text "p"}, if it succeeds; otherwise it returns 
 | 
| 104 | 167 |   the default value @{text "x"}. For example:
 | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 168 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 169 | @{ML_response [display,gray]
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 170 | "let | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 171 | val p = Scan.optional ($$ \"h\") \"x\" | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 172 | val input1 = Symbol.explode \"hello\" | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 173 | val input2 = Symbol.explode \"world\" | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 174 | in | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 175 | (p input1, p input2) | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 176 | end" | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 177 | "((\"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 | 178 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 179 |   The function @{ML_ind option in Scan} works similarly, except no default value can
 | 
| 50 | 180 |   be given. Instead, the result is wrapped as an @{text "option"}-type. For example:
 | 
| 181 | ||
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 182 | @{ML_response [display,gray]
 | 
| 50 | 183 | "let | 
| 184 | val p = Scan.option ($$ \"h\") | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 185 | val input1 = Symbol.explode \"hello\" | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 186 | val input2 = Symbol.explode \"world\" | 
| 50 | 187 | in | 
| 188 | (p input1, p input2) | |
| 189 | end" "((SOME \"h\", [\"e\", \"l\", \"l\", \"o\"]), (NONE, [\"w\", \"o\", \"r\", \"l\", \"d\"]))"} | |
| 49 | 190 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 191 |   The function @{ML_ind ahead in Scan} parses some input, but leaves the original
 | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 192 | input unchanged. For example: | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 193 | |
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 194 |   @{ML_response [display,gray]
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 195 | "Scan.ahead (Scan.this_string \"foo\") (Symbol.explode \"foo\")" | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 196 | "(\"foo\", [\"f\", \"o\", \"o\"])"} | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 197 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 198 |   The function @{ML_ind "!!" in Scan} helps with producing appropriate error messages
 | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 199 |   during parsing. For example if you want to parse @{text p} immediately 
 | 
| 58 | 200 |   followed by @{text q}, or start a completely different parser @{text r},
 | 
| 104 | 201 | you might write: | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 202 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 203 |   @{ML [display,gray] "(p -- q) || r" for p q r}
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 204 | |
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 205 | However, this parser is problematic for producing a useful error | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 206 |   message, if the parsing of @{ML "(p -- q)" for p q} fails. Because with the
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 207 |   parser above you lose the information that @{text p} should be followed by @{text q}.
 | 
| 220 | 208 |   To see this assume that @{text p} is present in the input, but it is not
 | 
| 209 |   followed by @{text q}. That means @{ML "(p -- q)" for p q} will fail and
 | |
| 210 |   hence the alternative parser @{text r} will be tried. However, in many
 | |
| 236 | 211 |   circumstances this will be the wrong parser for the input ``@{text "p"}-followed-by-something''
 | 
| 220 | 212 | and therefore will also fail. The error message is then caused by the failure | 
| 213 |   of @{text r}, not by the absence of @{text q} in the input. This kind of
 | |
| 214 |   situation can be avoided when using the function @{ML "!!"}.  This function
 | |
| 215 | aborts the whole process of parsing in case of a failure and prints an error | |
| 216 | message. For example if you invoke the parser | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 217 | |
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 218 | |
| 236 | 219 |   @{ML [display,gray] "!! (fn _ => \"foo\") ($$ \"h\")"}
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 220 | |
| 58 | 221 |   on @{text [quotes] "hello"}, the parsing succeeds
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 222 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 223 |   @{ML_response [display,gray] 
 | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 224 | "(!! (fn _ => \"foo\") ($$ \"h\")) (Symbol.explode \"hello\")" | 
| 236 | 225 | "(\"h\", [\"e\", \"l\", \"l\", \"o\"])"} | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 226 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 227 |   but if you invoke it on @{text [quotes] "world"}
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 228 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 229 |   @{ML_response_fake [display,gray] "(!! (fn _ => \"foo\") ($$ \"h\")) (Symbol.explode \"world\")"
 | 
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 230 | "Exception ABORT raised"} | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 231 | |
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
105diff
changeset | 232 |   then the parsing aborts and the error message @{text "foo"} is printed. In order to
 | 
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: 
116diff
changeset | 233 | see the error message properly, you need to prefix the parser with the function | 
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 234 |   @{ML_ind error in Scan}. For example:
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 235 | |
| 236 | 236 |   @{ML_response_fake [display,gray] 
 | 
| 237 | "Scan.error (!! (fn _ => \"foo\") ($$ \"h\"))" | |
| 238 | "Exception Error \"foo\" raised"} | |
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 239 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 240 |   This ``prefixing'' is usually done by wrappers such as @{ML_ind local_theory in OuterSyntax} 
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 241 |   (see Section~\ref{sec:newcommand} which explains this function in more detail). 
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 242 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 243 |   Let us now return to our example of parsing @{ML "(p -- q) || r" for p q
 | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 244 | r}. If you want to generate the correct error message for failure | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 245 |   of parsing @{text "p"}-followed-by-@{text "q"}, then you have to write:
 | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 246 | *} | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 247 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 248 | ML{*fun p_followed_by_q p q r =
 | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 249 | let | 
| 236 | 250 | val err_msg = fn _ => p ^ " is not followed by " ^ q | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 251 | in | 
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 252 | ($$ p -- (!! err_msg ($$ q))) || ($$ r -- $$ r) | 
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 253 | end *} | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 254 | |
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 255 | |
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 256 | text {*
 | 
| 220 | 257 | Running this parser with the arguments | 
| 258 |   @{text [quotes] "h"}, @{text [quotes] "e"} and @{text [quotes] "w"}, and 
 | |
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 259 |   the input @{text [quotes] "holle"} 
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 260 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 261 |   @{ML_response_fake [display,gray] "Scan.error (p_followed_by_q \"h\" \"e\" \"w\") (Symbol.explode \"holle\")"
 | 
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 262 | "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 | 263 | |
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 264 | produces the correct error message. Running it with | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 265 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 266 |   @{ML_response [display,gray] "Scan.error (p_followed_by_q \"h\" \"e\" \"w\") (Symbol.explode \"wworld\")"
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 267 | "((\"w\", \"w\"), [\"o\", \"r\", \"l\", \"d\"])"} | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 268 | |
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 269 | yields the expected parsing. | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 270 | |
| 58 | 271 |   The function @{ML "Scan.repeat p" for p} will apply a parser @{text p} as 
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 272 | often as it succeeds. For example: | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 273 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 274 |   @{ML_response [display,gray] "Scan.repeat ($$ \"h\") (Symbol.explode \"hhhhello\")" 
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 275 | "([\"h\", \"h\", \"h\", \"h\"], [\"e\", \"l\", \"l\", \"o\"])"} | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 276 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 277 |   Note that @{ML_ind repeat in Scan} stores the parsed items in a list. The function
 | 
| 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 278 |   @{ML_ind repeat1 in Scan} is similar, but requires that the parser @{text "p"} 
 | 
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 279 | succeeds at least once. | 
| 48 
609f9ef73494
fixed FIXME's in fake responses
 Christian Urban <urbanc@in.tum.de> parents: 
47diff
changeset | 280 | |
| 58 | 281 |   Also note that the parser would have aborted with the exception @{text MORE}, if
 | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 282 |   you had it run with the string @{text [quotes] "hhhh"}. This can be avoided by using
 | 
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 283 |   the wrapper @{ML_ind finite in Scan} and the ``stopper-token'' 
 | 
| 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 284 |   @{ML_ind stopper in Symbol}. With them you can write:
 | 
| 49 | 285 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 286 |   @{ML_response [display,gray] "Scan.finite Symbol.stopper (Scan.repeat ($$ \"h\")) (Symbol.explode \"hhhh\")" 
 | 
| 49 | 287 | "([\"h\", \"h\", \"h\", \"h\"], [])"} | 
| 288 | ||
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 289 |   The function @{ML stopper in Symbol} is the ``end-of-input'' indicator for parsing strings;
 | 
| 128 | 290 | other stoppers need to be used when parsing, for example, tokens. However, this kind of | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 291 | manually wrapping is often already done by the surrounding infrastructure. | 
| 49 | 292 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 293 |   The function @{ML_ind repeat in Scan} can be used with @{ML_ind one in Scan} to read any 
 | 
| 56 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 294 | 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 | 295 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 296 |   @{ML_response [display,gray] 
 | 
| 56 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 297 | "let | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 298 | val p = Scan.repeat (Scan.one Symbol.not_eof) | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 299 | val input = Symbol.explode \"foo bar foo\" | 
| 56 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 300 | in | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 301 | 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 | 302 | end" | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 303 | "([\"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 | 304 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 305 |   where the function @{ML_ind not_eof in Symbol} ensures that we do not read beyond the 
 | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 306 | end of the input string (i.e.~stopper symbol). | 
| 56 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 307 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 308 |   The function @{ML_ind unless in Scan} takes two parsers: if the first one can 
 | 
| 60 
5b9c6010897b
doem tuning and made the cookbook work again with recent changes (CookBook/Package/Ind_Interface.thy needs to be looked at to see what the problem with the new parser type is)
 Christian Urban <urbanc@in.tum.de> parents: 
59diff
changeset | 309 | parse the input, then the whole parser fails; if not, then the second is tried. Therefore | 
| 56 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 310 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 311 |   @{ML_response_fake_both [display,gray] "Scan.unless ($$ \"h\") ($$ \"w\") (Symbol.explode \"hello\")"
 | 
| 56 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 312 | "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 | 313 | |
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 314 | fails, while | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 315 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 316 |   @{ML_response [display,gray] "Scan.unless ($$ \"h\") ($$ \"w\") (Symbol.explode \"world\")"
 | 
| 56 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 317 | "(\"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 | 318 | |
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 319 | succeeds. | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 320 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 321 |   The functions @{ML_ind repeat in Scan} and @{ML_ind unless in Scan} can 
 | 
| 256 
1fb8d62c88a0
added some first index-information
 Christian Urban <urbanc@in.tum.de> parents: 
250diff
changeset | 322 | be combined to read any input until a certain marker symbol is reached. In the | 
| 
1fb8d62c88a0
added some first index-information
 Christian Urban <urbanc@in.tum.de> parents: 
250diff
changeset | 323 |   example below the marker symbol is a @{text [quotes] "*"}.
 | 
| 56 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 324 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 325 |   @{ML_response [display,gray]
 | 
| 56 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 326 | "let | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 327 | val p = Scan.repeat (Scan.unless ($$ \"*\") (Scan.one Symbol.not_eof)) | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 328 | val input1 = Symbol.explode \"fooooo\" | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 329 | val input2 = Symbol.explode \"foo*ooo\" | 
| 56 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 330 | in | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 331 | (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 | 332 | 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 | 333 | end" | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 334 | "(([\"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 | 335 | ([\"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 | 336 | |
| 256 
1fb8d62c88a0
added some first index-information
 Christian Urban <urbanc@in.tum.de> parents: 
250diff
changeset | 337 | |
| 220 | 338 | After parsing is done, you almost always want to apply a function to the parsed | 
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 339 |   items. One way to do this is the function @{ML_ind ">>" in Scan} where 
 | 
| 256 
1fb8d62c88a0
added some first index-information
 Christian Urban <urbanc@in.tum.de> parents: 
250diff
changeset | 340 |   @{ML "(p >> f)" for p f} runs 
 | 
| 58 | 341 |   first the parser @{text p} and upon successful completion applies the 
 | 
| 342 |   function @{text f} to the result. For example
 | |
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 343 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 344 | @{ML_response [display,gray]
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 345 | "let | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 346 | fun double (x, y) = (x ^ x, y ^ y) | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 347 | val parser = $$ \"h\" -- $$ \"e\" | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 348 | in | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 349 | (parser >> double) (Symbol.explode \"hello\") | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 350 | end" | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 351 | "((\"hh\", \"ee\"), [\"l\", \"l\", \"o\"])"} | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 352 | |
| 104 | 353 | doubles the two parsed input strings; or | 
| 59 | 354 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 355 |   @{ML_response [display,gray] 
 | 
| 59 | 356 | "let | 
| 104 | 357 | val p = Scan.repeat (Scan.one Symbol.not_eof) | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 358 | val input = Symbol.explode \"foo bar foo\" | 
| 59 | 359 | in | 
| 104 | 360 | Scan.finite Symbol.stopper (p >> implode) input | 
| 59 | 361 | end" | 
| 362 | "(\"foo bar foo\",[])"} | |
| 363 | ||
| 60 
5b9c6010897b
doem tuning and made the cookbook work again with recent changes (CookBook/Package/Ind_Interface.thy needs to be looked at to see what the problem with the new parser type is)
 Christian Urban <urbanc@in.tum.de> parents: 
59diff
changeset | 364 | where the single-character strings in the parsed output are transformed | 
| 59 | 365 | back into one string. | 
| 56 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 366 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 367 |   The function @{ML_ind lift in Scan} takes a parser and a pair as arguments. This function applies
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 368 | 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 | 369 | untouched. For example | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 370 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 371 | @{ML_response [display,gray]
 | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 372 | "Scan.lift ($$ \"h\" -- $$ \"e\") (1, Symbol.explode \"hello\")" | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 373 | "((\"h\", \"e\"), (1, [\"l\", \"l\", \"o\"]))"} | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 374 | |
| 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 | 375 | (FIXME: In which situations is this useful? Give examples.) | 
| 149 | 376 | |
| 377 |   \begin{exercise}\label{ex:scancmts}
 | |
| 378 | Write a parser that parses an input string so that any comment enclosed | |
| 220 | 379 |   within @{text "(*\<dots>*)"} is replaced by the same comment but enclosed within
 | 
| 149 | 380 |   @{text "(**\<dots>**)"} in the output string. To enclose a string, you can use the
 | 
| 381 |   function @{ML "enclose s1 s2 s" for s1 s2 s} which produces the string @{ML
 | |
| 236 | 382 | "s1 ^ s ^ s2" for s1 s2 s}. Hint: To simplify the task ignore the proper | 
| 383 | nesting of comments. | |
| 149 | 384 |   \end{exercise}
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 385 | *} | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 386 | |
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 387 | section {* Parsing Theory Syntax *}
 | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 388 | |
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 389 | text {*
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 390 | Most of the time, however, Isabelle developers have to deal with parsing | 
| 156 | 391 | tokens, not strings. These token parsers have the type: | 
| 128 | 392 | *} | 
| 393 | ||
| 394 | ML{*type 'a parser = OuterLex.token list -> 'a * OuterLex.token list*}
 | |
| 395 | ||
| 396 | text {*
 | |
| 149 | 397 | The reason for using token parsers is that theory syntax, as well as the | 
| 128 | 398 |   parsers for the arguments of proof methods, use the type @{ML_type
 | 
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 399 | OuterLex.token}. | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 400 | |
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 401 |   \begin{readmore}
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 402 | 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 | 403 |   @{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 | 404 |   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 | 405 |   \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 | 406 | |
| 316 
74f0a06f751f
further polishing of index generation
 Christian Urban <urbanc@in.tum.de> parents: 
315diff
changeset | 407 |   The structure @{ML_struct  OuterLex} defines several kinds of tokens (for
 | 
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 408 |   example @{ML_ind Ident in OuterLex} for identifiers, @{ML Keyword in
 | 
| 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 409 |   OuterLex} for keywords and @{ML_ind Command in OuterLex} for commands). Some
 | 
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 410 | token parsers take into account the kind of tokens. The first example shows | 
| 256 
1fb8d62c88a0
added some first index-information
 Christian Urban <urbanc@in.tum.de> parents: 
250diff
changeset | 411 | how to generate a token list out of a string using the function | 
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 412 |   @{ML_ind scan in OuterSyntax}. It is given the argument 
 | 
| 256 
1fb8d62c88a0
added some first index-information
 Christian Urban <urbanc@in.tum.de> parents: 
250diff
changeset | 413 |   @{ML "Position.none"} since,
 | 
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 414 | at the moment, we are not interested in generating precise error | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 415 |   messages. The following code\footnote{Note that because of a possible bug in
 | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 416 |   the PolyML runtime system, the result is printed as @{text [quotes] "?"},
 | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 417 | instead of the tokens.} | 
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 418 | |
| 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 | 419 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 420 | @{ML_response_fake [display,gray] "OuterSyntax.scan Position.none \"hello world\"" 
 | 
| 50 | 421 | "[Token (\<dots>,(Ident, \"hello\"),\<dots>), | 
| 422 | Token (\<dots>,(Space, \" \"),\<dots>), | |
| 423 | Token (\<dots>,(Ident, \"world\"),\<dots>)]"} | |
| 424 | ||
| 425 | produces three tokens where the first and the last are identifiers, since | |
| 58 | 426 |   @{text [quotes] "hello"} and @{text [quotes] "world"} do not match any
 | 
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 427 | other syntactic category. The second indicates a space. | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 428 | |
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 429 | We can easily change what is recognised as a keyword with the function | 
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 430 |   @{ML_ind keyword in OuterKeyword}. For example calling it with 
 | 
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 431 | *} | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 432 | |
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 433 | ML{*val _ = OuterKeyword.keyword "hello"*}
 | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 434 | |
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 435 | text {*
 | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 436 |   then lexing @{text [quotes] "hello world"} will produce
 | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 437 | |
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 438 |   @{ML_response_fake [display,gray] "OuterSyntax.scan Position.none \"hello world\"" 
 | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 439 | "[Token (\<dots>,(Keyword, \"hello\"),\<dots>), | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 440 | Token (\<dots>,(Space, \" \"),\<dots>), | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 441 | Token (\<dots>,(Ident, \"world\"),\<dots>)]"} | 
| 50 | 442 | |
| 241 | 443 | Many parsing functions later on will require white space, 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 | 444 | to have already been filtered out. So from now on we are going to use the | 
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 445 |   functions @{ML filter} and @{ML_ind is_proper in OuterLex} to do this. 
 | 
| 256 
1fb8d62c88a0
added some first index-information
 Christian Urban <urbanc@in.tum.de> parents: 
250diff
changeset | 446 | 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 | 447 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 448 | @{ML_response_fake [display,gray]
 | 
| 50 | 449 | "let | 
| 450 | val input = OuterSyntax.scan Position.none \"hello world\" | |
| 451 | in | |
| 452 | filter OuterLex.is_proper input | |
| 453 | end" | |
| 454 | "[Token (\<dots>,(Ident, \"hello\"), \<dots>), Token (\<dots>,(Ident, \"world\"), \<dots>)]"} | |
| 455 | ||
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 456 | For convenience we define the function: | 
| 50 | 457 | *} | 
| 458 | ||
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 459 | ML{*fun filtered_input str = 
 | 
| 160 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 460 | filter OuterLex.is_proper (OuterSyntax.scan Position.none str) *} | 
| 50 | 461 | |
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 462 | text {* 
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 463 | If you now 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 | 464 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 465 | @{ML_response_fake [display,gray] 
 | 
| 50 | 466 | "filtered_input \"inductive | for\"" | 
| 467 | "[Token (\<dots>,(Command, \"inductive\"),\<dots>), | |
| 468 | Token (\<dots>,(Keyword, \"|\"),\<dots>), | |
| 469 | 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 | 470 | |
| 221 | 471 | you obtain a list consisting of only one command and two keyword tokens. | 
| 241 | 472 | If you want to see which keywords and commands are currently known to Isabelle, | 
| 473 | type: | |
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 474 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 475 | @{ML_response_fake [display,gray] 
 | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 476 | "let | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 477 | 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 | 478 | in | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 479 | (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 | 480 | end" | 
| 132 | 481 | "([\"}\", \"{\", \<dots>], [\"\<rightleftharpoons>\", \"\<leftharpoondown>\", \<dots>])"}
 | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 482 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 483 |   You might have to adjust the @{ML_ind print_depth} in order to
 | 
| 241 | 484 | see the complete list. | 
| 485 | ||
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 486 |   The parser @{ML_ind "$$$" in OuterParse} parses a single keyword. For example:
 | 
| 50 | 487 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 488 | @{ML_response [display,gray]
 | 
| 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 | 489 | "let | 
| 50 | 490 | val input1 = filtered_input \"where for\" | 
| 491 | 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 | 492 | 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 | 493 | (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 | 494 | end" | 
| 128 | 495 | "((\"where\",\<dots>), (\"|\",\<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 | 496 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 497 |   Any non-keyword string can be parsed with the function @{ML_ind reserved in OuterParse}.
 | 
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 498 | For example: | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 499 | |
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 500 |   @{ML_response [display,gray]
 | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 501 | "let | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 502 | val p = OuterParse.reserved \"bar\" | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 503 | val input = filtered_input \"bar\" | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 504 | in | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 505 | p input | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 506 | end" | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 507 | "(\"bar\",[])"} | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 508 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 509 |   Like before, you can sequentially connect parsers with @{ML "--"}. 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 | 510 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 511 | @{ML_response [display,gray]
 | 
| 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 | "let | 
| 50 | 513 | 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 | 514 | 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 | 515 | (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 | 516 | end" | 
| 183 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
181diff
changeset | 517 | "((\"|\", \"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 | 518 | |
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 519 |   The parser @{ML "OuterParse.enum s p" for s p} parses a possibly empty 
 | 
| 58 | 520 |   list of items recognised by the parser @{text p}, where the items being parsed
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 521 |   are separated by the string @{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 | 522 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 523 | @{ML_response [display,gray]
 | 
| 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 | 524 | "let | 
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 525 | 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 | 526 | 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 | 527 | (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 | 528 | end" | 
| 183 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
181diff
changeset | 529 | "([\"in\", \"in\", \"in\"], [\<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 | 530 | |
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 531 |   The function @{ML_ind enum1 in OuterParse} works similarly, except that the
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 532 |   parsed list must be non-empty. Note that we had to add a string @{text
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 533 | [quotes] "foo"} at the end of the parsed string, otherwise the parser would | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 534 |   have consumed all tokens and then failed with the exception @{text
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 535 | "MORE"}. Like in the previous section, we can avoid this exception using the | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 536 |   wrapper @{ML Scan.finite}. This time, however, we have to use the
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 537 |   ``stopper-token'' @{ML OuterLex.stopper}. We can write:
 | 
| 49 | 538 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 539 | @{ML_response [display,gray]
 | 
| 49 | 540 | "let | 
| 50 | 541 | val input = filtered_input \"in | in | in\" | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 542 | val p = OuterParse.enum \"|\" (OuterParse.$$$ \"in\") | 
| 49 | 543 | in | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 544 | Scan.finite OuterLex.stopper p input | 
| 49 | 545 | end" | 
| 183 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
181diff
changeset | 546 | "([\"in\", \"in\", \"in\"], [])"} | 
| 49 | 547 | |
| 75 | 548 | The following function will help to run examples. | 
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 549 | *} | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 550 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 551 | ML{*fun parse p input = Scan.finite OuterLex.stopper (Scan.error p) input *}
 | 
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 552 | |
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 553 | text {*
 | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 554 |   The function @{ML_ind "!!!" in OuterParse} can be used to force termination
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 555 |   of the parser in case of a dead end, just like @{ML "Scan.!!"} (see previous
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 556 |   section). A difference, however, is that the error message of @{ML
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 557 |   "OuterParse.!!!"} is fixed to be @{text [quotes] "Outer syntax error"}
 | 
| 221 | 558 | together with a relatively precise description of the failure. For example: | 
| 49 | 559 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 560 | @{ML_response_fake [display,gray]
 | 
| 49 | 561 | "let | 
| 50 | 562 | val input = filtered_input \"in |\" | 
| 49 | 563 | val parse_bar_then_in = OuterParse.$$$ \"|\" -- OuterParse.$$$ \"in\" | 
| 564 | in | |
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 565 | parse (OuterParse.!!! parse_bar_then_in) input | 
| 49 | 566 | end" | 
| 567 | "Exception ERROR \"Outer syntax error: keyword \"|\" expected, | |
| 568 | but keyword in was found\" raised" | |
| 569 | } | |
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 570 | |
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 571 |   \begin{exercise} (FIXME)
 | 
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 572 |   A type-identifier, for example @{typ "'a"}, is a token of 
 | 
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 573 |   kind @{ML_ind Keyword in OuterLex}. It can be parsed using 
 | 
| 256 
1fb8d62c88a0
added some first index-information
 Christian Urban <urbanc@in.tum.de> parents: 
250diff
changeset | 574 |   the function @{ML type_ident in OuterParse}.
 | 
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 575 |   \end{exercise}
 | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 576 | |
| 104 | 577 | (FIXME: or give parser for numbers) | 
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 578 | |
| 125 | 579 | Whenever there is a possibility that the processing of user input can fail, | 
| 221 | 580 | it is a good idea to give all available information about where the error | 
| 220 | 581 | occurred. For this Isabelle can attach positional information to tokens | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 582 | and then thread this information up the ``processing chain''. To see this, | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 583 |   modify the function @{ML filtered_input}, described earlier, as follows 
 | 
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 584 | *} | 
| 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 585 | |
| 125 | 586 | ML{*fun filtered_input' str = 
 | 
| 587 | filter OuterLex.is_proper (OuterSyntax.scan (Position.line 7) str) *} | |
| 49 | 588 | |
| 589 | text {*
 | |
| 125 | 590 | where we pretend the parsed string starts on line 7. An example is | 
| 49 | 591 | |
| 125 | 592 | @{ML_response_fake [display,gray]
 | 
| 593 | "filtered_input' \"foo \\n bar\"" | |
| 594 | "[Token ((\"foo\", ({line=7, end_line=7}, {line=7})), (Ident, \"foo\"), \<dots>),
 | |
| 595 |  Token ((\"bar\", ({line=8, end_line=8}, {line=8})), (Ident, \"bar\"), \<dots>)]"}
 | |
| 596 | ||
| 597 |   in which the @{text [quotes] "\\n"} causes the second token to be in 
 | |
| 598 | line 8. | |
| 599 | ||
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 600 |   By using the parser @{ML position in OuterParse} you can access the token 
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 601 | position and return it as part of the parser result. For example | 
| 125 | 602 | |
| 603 | @{ML_response_fake [display,gray]
 | |
| 604 | "let | |
| 241 | 605 | val input = filtered_input' \"where\" | 
| 125 | 606 | in | 
| 607 | parse (OuterParse.position (OuterParse.$$$ \"where\")) input | |
| 608 | end" | |
| 609 | "((\"where\", {line=7, end_line=7}), [])"}
 | |
| 610 | ||
| 611 |   \begin{readmore}
 | |
| 612 | The functions related to positions are implemented in the file | |
| 613 |   @{ML_file "Pure/General/position.ML"}.
 | |
| 614 |   \end{readmore}
 | |
| 49 | 615 | |
| 616 | *} | |
| 617 | ||
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 618 | section {* Parsers for ML-Code (TBD) *}
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 619 | |
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 620 | text {*
 | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 621 |   @{ML_ind ML_source in OuterParse}
 | 
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 622 | *} | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 623 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 624 | section {* Context Parser (TBD) *}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 625 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 626 | text {*
 | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 627 |   @{ML_ind Args.context}
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 628 | *} | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 629 | (* | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 630 | ML {*
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 631 | let | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 632 | val parser = Args.context -- Scan.lift Args.name_source | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 633 | |
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 634 | fun term_pat (ctxt, str) = | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 635 | str |> Syntax.read_prop ctxt | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 636 | in | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 637 |   (parser >> term_pat) (Context.Proof @{context}, filtered_input "f (a::nat)")
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 638 | |> fst | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 639 | end | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 640 | *} | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 641 | *) | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 642 | |
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 643 | text {*
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 644 |   @{ML_ind Args.context}
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 645 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 646 |   Used for example in \isacommand{attribute\_setup} and \isacommand{method\_setup}.
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 647 | *} | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 648 | |
| 207 | 649 | section {* Argument and Attribute Parsers (TBD) *}
 | 
| 650 | ||
| 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 | 651 | section {* Parsing Inner Syntax *}
 | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 652 | |
| 125 | 653 | text {*
 | 
| 654 | There is usually no need to write your own parser for parsing inner syntax, that is | |
| 285 | 655 | for terms and types: you can just call the predefined parsers. Terms can | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 656 |   be parsed using the function @{ML_ind term in OuterParse}. For example:
 | 
| 125 | 657 | |
| 658 | @{ML_response [display,gray]
 | |
| 659 | "let | |
| 660 | val input = OuterSyntax.scan Position.none \"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 | 661 | in | 
| 125 | 662 | OuterParse.term input | 
| 663 | end" | |
| 664 | "(\"\\^E\\^Ftoken\\^Efoo\\^E\\^F\\^E\", [])"} | |
| 665 | ||
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 666 |   The function @{ML_ind prop in OuterParse} is similar, except that it gives a different
 | 
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 667 | error message, when parsing fails. As you can see, the parser not just returns | 
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 668 | the parsed string, but also some encoded information. You can decode the | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 669 |   information with the function @{ML_ind parse in YXML} in @{ML_struct YXML}. For example
 | 
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 670 | |
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 671 |   @{ML_response [display,gray]
 | 
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 672 | "YXML.parse \"\\^E\\^Ftoken\\^Efoo\\^E\\^F\\^E\"" | 
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 673 | "XML.Elem (\"token\", [], [XML.Text \"foo\"])"} | 
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 674 | |
| 149 | 675 | The result of the decoding is an XML-tree. You can see better what is going on if | 
| 131 | 676 |   you replace @{ML Position.none} by @{ML "Position.line 42"}, say:
 | 
| 101 | 677 | |
| 125 | 678 | @{ML_response [display,gray]
 | 
| 679 | "let | |
| 680 | val input = OuterSyntax.scan (Position.line 42) \"foo\" | |
| 681 | in | |
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 682 | YXML.parse (fst (OuterParse.term input)) | 
| 125 | 683 | end" | 
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 684 | "XML.Elem (\"token\", [(\"line\", \"42\"), (\"end_line\", \"42\")], [XML.Text \"foo\"])"} | 
| 125 | 685 | |
| 149 | 686 | The positional information is stored as part of an XML-tree so that code | 
| 687 | called later on will be able to give more precise error messages. | |
| 125 | 688 | |
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 689 |   \begin{readmore}
 | 
| 128 | 690 | The functions to do with input and output of XML and YXML are defined | 
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 691 |   in @{ML_file "Pure/General/xml.ML"} and @{ML_file "Pure/General/yxml.ML"}.
 | 
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 692 |   \end{readmore}
 | 
| 160 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 693 | |
| 361 | 694 | FIXME: | 
| 695 |   @{ML_ind parse_term in Syntax} @{ML_ind check_term in Syntax}
 | |
| 696 |   @{ML_ind parse_typ in Syntax} @{ML_ind check_typ in Syntax}
 | |
| 697 | ||
| 698 | ||
| 125 | 699 | *} | 
| 101 | 700 | |
| 116 
c9ff326e3ce5
more changes to the package chapter
 Christian Urban <urbanc@in.tum.de> parents: 
114diff
changeset | 701 | section {* Parsing Specifications\label{sec:parsingspecs} *}
 | 
| 101 | 702 | |
| 703 | text {*
 | |
| 121 | 704 | There are a number of special purpose parsers that help with parsing | 
| 156 | 705 | specifications of function definitions, inductive predicates and so on. In | 
| 220 | 706 |   Chapter~\ref{chp:package}, for example, we will need to parse specifications
 | 
| 121 | 707 | for inductive predicates of the form: | 
| 708 | *} | |
| 101 | 709 | |
| 121 | 710 | simple_inductive | 
| 711 | even and odd | |
| 712 | where | |
| 713 | even0: "even 0" | |
| 714 | | evenS: "odd n \<Longrightarrow> even (Suc n)" | |
| 715 | | oddS: "even n \<Longrightarrow> odd (Suc n)" | |
| 101 | 716 | |
| 327 
ce754ad78bc9
more work on the storing section
 Christian Urban <urbanc@in.tum.de> parents: 
326diff
changeset | 717 | |
| 101 | 718 | text {*
 | 
| 121 | 719 | For this we are going to use the parser: | 
| 101 | 720 | *} | 
| 721 | ||
| 121 | 722 | ML %linenosgray{*val spec_parser = 
 | 
| 126 | 723 | OuterParse.fixes -- | 
| 724 | Scan.optional | |
| 725 | (OuterParse.$$$ "where" |-- | |
| 726 | OuterParse.!!! | |
| 727 | (OuterParse.enum1 "|" | |
| 728 | (SpecParse.opt_thm_name ":" -- OuterParse.prop))) []*} | |
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: 
116diff
changeset | 729 | |
| 101 | 730 | text {*
 | 
| 241 | 731 | Note that the parser must not parse the keyword \simpleinductive, even if it is | 
| 126 | 732 | meant to process definitions as shown above. The parser of the keyword | 
| 128 | 733 |   will be given by the infrastructure that will eventually call @{ML spec_parser}.
 | 
| 126 | 734 | |
| 735 | ||
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 736 | To see what the parser returns, let us parse the string corresponding to the | 
| 121 | 737 |   definition of @{term even} and @{term odd}:
 | 
| 738 | ||
| 101 | 739 | @{ML_response [display,gray]
 | 
| 740 | "let | |
| 741 | val input = filtered_input | |
| 742 | (\"even and odd \" ^ | |
| 743 | \"where \" ^ | |
| 744 | \" even0[intro]: \\\"even 0\\\" \" ^ | |
| 745 | \"| evenS[intro]: \\\"odd n \<Longrightarrow> even (Suc n)\\\" \" ^ | |
| 746 | \"| oddS[intro]: \\\"even n \<Longrightarrow> odd (Suc n)\\\"\") | |
| 747 | in | |
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: 
116diff
changeset | 748 | parse spec_parser input | 
| 101 | 749 | end" | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
183diff
changeset | 750 | "(([(even, NONE, NoSyn), (odd, NONE, NoSyn)], | 
| 101 | 751 | [((even0,\<dots>), \"\\^E\\^Ftoken\\^Eeven 0\\^E\\^F\\^E\"), | 
| 752 | ((evenS,\<dots>), \"\\^E\\^Ftoken\\^Eodd n \<Longrightarrow> even (Suc n)\\^E\\^F\\^E\"), | |
| 753 | ((oddS,\<dots>), \"\\^E\\^Ftoken\\^Eeven n \<Longrightarrow> odd (Suc n)\\^E\\^F\\^E\")]), [])"} | |
| 121 | 754 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
183diff
changeset | 755 | As you see, the result is a pair consisting of a list of | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
183diff
changeset | 756 | variables with optional type-annotation and syntax-annotation, and a list of | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
183diff
changeset | 757 | rules where every rule has optionally a name and an attribute. | 
| 121 | 758 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 759 |   The function @{ML_ind "fixes" in OuterParse} in Line 2 of the parser reads an 
 | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
183diff
changeset | 760 |   \isacommand{and}-separated 
 | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 761 | list of variables that can include optional type annotations and syntax translations. | 
| 121 | 762 |   For example:\footnote{Note that in the code we need to write 
 | 
| 763 |   @{text "\\\"int \<Rightarrow> bool\\\""} in order to properly escape the double quotes
 | |
| 764 | in the compound type.} | |
| 765 | ||
| 766 | @{ML_response [display,gray]
 | |
| 767 | "let | |
| 768 | val input = filtered_input | |
| 769 | \"foo::\\\"int \<Rightarrow> bool\\\" and bar::nat (\\\"BAR\\\" 100) and blonk\" | |
| 770 | in | |
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 771 | parse OuterParse.fixes input | 
| 121 | 772 | end" | 
| 773 | "([(foo, SOME \"\\^E\\^Ftoken\\^Eint \<Rightarrow> bool\\^E\\^F\\^E\", NoSyn), | |
| 774 | (bar, SOME \"\\^E\\^Ftoken\\^Enat\\^E\\^F\\^E\", Mixfix (\"BAR\", [], 100)), | |
| 775 | (blonk, NONE, NoSyn)],[])"} | |
| 50 | 776 | *} | 
| 777 | ||
| 121 | 778 | text {*
 | 
| 156 | 779 |   Whenever types are given, they are stored in the @{ML SOME}s. The types are
 | 
| 780 | not yet used to type the variables: this must be done by type-inference later | |
| 149 | 781 | on. Since types are part of the inner syntax they are strings with some | 
| 241 | 782 | encoded information (see previous section). If a mixfix-syntax is | 
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 783 |   present for a variable, then it is stored in the @{ML_ind Mixfix} data structure;
 | 
| 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 784 |   no syntax translation is indicated by @{ML_ind NoSyn}.
 | 
| 121 | 785 | |
| 786 |   \begin{readmore}
 | |
| 241 | 787 |   The data structure for mixfix annotations is defined in @{ML_file "Pure/Syntax/mixfix.ML"}.
 | 
| 121 | 788 |   \end{readmore}
 | 
| 789 | ||
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
183diff
changeset | 790 |   Lines 3 to 7 in the function @{ML spec_parser} implement the parser for a
 | 
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 791 | list of introduction rules, that is propositions with theorem annotations | 
| 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 792 | such as rule names and attributes. The introduction rules are propositions | 
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 793 |   parsed by @{ML_ind prop in OuterParse}. However, they can include an optional
 | 
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 794 | theorem name plus some attributes. For example | 
| 121 | 795 | |
| 796 | @{ML_response [display,gray] "let 
 | |
| 797 | val input = filtered_input \"foo_lemma[intro,dest!]:\" | |
| 798 | val ((name, attrib), _) = parse (SpecParse.thm_name \":\") input | |
| 799 | in | |
| 800 | (name, map Args.dest_src attrib) | |
| 801 | end" "(foo_lemma, [((\"intro\", []), \<dots>), ((\"dest\", [\<dots>]), \<dots>)])"} | |
| 802 | ||
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 803 |   The function @{ML_ind opt_thm_name in SpecParse} is the ``optional'' variant of
 | 
| 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 804 |   @{ML_ind thm_name in SpecParse}. Theorem names can contain attributes. The name 
 | 
| 131 | 805 |   has to end with @{text [quotes] ":"}---see the argument of 
 | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
183diff
changeset | 806 |   the function @{ML SpecParse.opt_thm_name} in Line 7.
 | 
| 121 | 807 | |
| 808 |   \begin{readmore}
 | |
| 809 |   Attributes and arguments are implemented in the files @{ML_file "Pure/Isar/attrib.ML"} 
 | |
| 810 |   and @{ML_file "Pure/Isar/args.ML"}.
 | |
| 811 |   \end{readmore}
 | |
| 101 | 812 | *} | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 813 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 814 | text_raw {*
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 815 |   \begin{exercise}
 | 
| 207 | 816 |   Have a look at how the parser @{ML SpecParse.where_alt_specs} is implemented
 | 
| 817 |   in file @{ML_file "Pure/Isar/spec_parse.ML"}. This parser corresponds
 | |
| 818 | to the ``where-part'' of the introduction rules given above. Below | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 819 |   we paraphrase the code of @{ML_ind where_alt_specs in SpecParse} adapted to our
 | 
| 207 | 820 | purposes. | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 821 |   \begin{isabelle}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 822 | *} | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 823 | ML %linenosgray{*val spec_parser' = 
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 824 | OuterParse.fixes -- | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 825 | Scan.optional | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 826 | (OuterParse.$$$ "where" |-- | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 827 | OuterParse.!!! | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 828 | (OuterParse.enum1 "|" | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 829 | ((SpecParse.opt_thm_name ":" -- OuterParse.prop) --| | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 830 | Scan.option (Scan.ahead (OuterParse.name || | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 831 | OuterParse.$$$ "[") -- | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 832 | OuterParse.!!! (OuterParse.$$$ "|"))))) [] *} | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 833 | text_raw {*
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 834 |   \end{isabelle}
 | 
| 284 | 835 | Both parsers accept the same input% that's not true: | 
| 836 | % spec_parser accepts input that is refuted by spec_parser' | |
| 837 | , but if you look closely, you can notice | |
| 207 | 838 |   an additional  ``tail'' (Lines 8 to 10) in @{ML spec_parser'}. What is the purpose of 
 | 
| 839 | this additional ``tail''? | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 840 |   \end{exercise}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 841 | *} | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 842 | |
| 229 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 843 | text {*
 | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 844 |   (FIXME: @{ML OuterParse.type_args}, @{ML OuterParse.typ}, @{ML OuterParse.opt_mixfix})
 | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 845 | *} | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 846 | |
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 847 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 848 | section {* New Commands and Keyword Files\label{sec:newcommand} *}
 | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 849 | |
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 850 | text {*
 | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 851 | Often new commands, for example for providing new definitional principles, | 
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 852 | need to be implemented. While this is not difficult on the ML-level, | 
| 66 | 853 | new commands, in order to be useful, need to be recognised by | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 854 | ProofGeneral. This results in some subtle configuration issues, which we | 
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 855 | will explain in this section. | 
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 856 | |
| 74 | 857 | To keep things simple, let us start with a ``silly'' command that does nothing | 
| 858 |   at all. We shall name this command \isacommand{foobar}. On the ML-level it can be 
 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 859 | defined as: | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 860 | *} | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 861 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 862 | ML{*let
 | 
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 863 | val do_nothing = Scan.succeed (LocalTheory.theory I) | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 864 | val kind = OuterKeyword.thy_decl | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 865 | in | 
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 866 | OuterSyntax.local_theory "foobar" "description of foobar" kind do_nothing | 
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 867 | end *} | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 868 | |
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 869 | text {*
 | 
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 870 |   The crucial function @{ML_ind local_theory in OuterSyntax} expects a name for the command, a
 | 
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 871 | short description, a kind indicator (which we will explain later more thoroughly) and a | 
| 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 872 | parser producing a local theory transition (its purpose will also explained | 
| 66 | 873 | later). | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 874 | |
| 101 | 875 | While this is everything you have to do on the ML-level, you need a keyword | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 876 | file that can be loaded by ProofGeneral. This is to enable ProofGeneral to | 
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 877 |   recognise \isacommand{foobar} as a command. Such a keyword file can be
 | 
| 74 | 878 | generated with the command-line: | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 879 | |
| 74 | 880 |   @{text [display] "$ isabelle keywords -k foobar some_log_files"}
 | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 881 | |
| 74 | 882 |   The option @{text "-k foobar"} indicates which postfix the name of the keyword file 
 | 
| 80 | 883 |   will be assigned. In the case above the file will be named @{text
 | 
| 86 | 884 | "isar-keywords-foobar.el"}. This command requires log files to be | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 885 | present (in order to extract the keywords from them). To generate these log | 
| 101 | 886 | files, you first need to package the code above into a separate theory file named | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 887 |   @{text "Command.thy"}, say---see Figure~\ref{fig:commandtheory} for the
 | 
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 888 | complete code. | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 889 | |
| 66 | 890 | |
| 891 | %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% | |
| 892 |   \begin{figure}[t]
 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 893 |   \begin{graybox}\small
 | 
| 66 | 894 |   \isacommand{theory}~@{text Command}\\
 | 
| 895 |   \isacommand{imports}~@{text Main}\\
 | |
| 896 |   \isacommand{begin}\\
 | |
| 85 
b02904872d6b
better handling of {* and *}
 Christian Urban <urbanc@in.tum.de> parents: 
81diff
changeset | 897 |   \isacommand{ML}~@{text "\<verbopen>"}\\
 | 
| 66 | 898 |   @{ML
 | 
| 899 | "let | |
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 900 | val do_nothing = Scan.succeed (LocalTheory.theory I) | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 901 | val kind = OuterKeyword.thy_decl | 
| 66 | 902 | in | 
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 903 | OuterSyntax.local_theory \"foobar\" \"description of foobar\" kind do_nothing | 
| 66 | 904 | end"}\\ | 
| 85 
b02904872d6b
better handling of {* and *}
 Christian Urban <urbanc@in.tum.de> parents: 
81diff
changeset | 905 |   @{text "\<verbclose>"}\\
 | 
| 66 | 906 |   \isacommand{end}
 | 
| 80 | 907 |   \end{graybox}
 | 
| 241 | 908 |   \caption{This file can be used to generate a log file. This log file in turn can
 | 
| 909 |   be used to generate a keyword file containing the command \isacommand{foobar}.
 | |
| 910 |   \label{fig:commandtheory}}
 | |
| 66 | 911 |   \end{figure}
 | 
| 912 | %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% | |
| 913 | ||
| 75 | 914 | For our purposes it is sufficient to use the log files of the theories | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 915 |   @{text "Pure"}, @{text "HOL"} and @{text "Pure-ProofGeneral"}, as well as
 | 
| 75 | 916 |   the log file for the theory @{text "Command.thy"}, which contains the new
 | 
| 917 |   \isacommand{foobar}-command. If you target other logics besides HOL, such
 | |
| 74 | 918 | as Nominal or ZF, then you need to adapt the log files appropriately. | 
| 104 | 919 | |
| 74 | 920 |   @{text Pure} and @{text HOL} are usually compiled during the installation of
 | 
| 921 | Isabelle. So log files for them should be already available. If not, then | |
| 75 | 922 | they can be conveniently compiled with the help of the build-script from the Isabelle | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 923 | distribution. | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 924 | |
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 925 |   @{text [display] 
 | 
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 926 | "$ ./build -m \"Pure\" | 
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 927 | $ ./build -m \"HOL\""} | 
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 928 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 929 |   The @{text "Pure-ProofGeneral"} theory needs to be compiled with:
 | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 930 | |
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 931 |   @{text [display] "$ ./build -m \"Pure-ProofGeneral\" \"Pure\""}
 | 
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 932 | |
| 101 | 933 |   For the theory @{text "Command.thy"}, you first need to create a ``managed'' subdirectory 
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 934 | with: | 
| 66 | 935 | |
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 936 |   @{text [display] "$ isabelle mkdir FoobarCommand"}
 | 
| 66 | 937 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 938 | This generates a directory containing the files: | 
| 66 | 939 | |
| 940 |   @{text [display] 
 | |
| 941 | "./IsaMakefile | |
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 942 | ./FoobarCommand/ROOT.ML | 
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 943 | ./FoobarCommand/document | 
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 944 | ./FoobarCommand/document/root.tex"} | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 945 | |
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 946 | |
| 101 | 947 |   You need to copy the file @{text "Command.thy"} into the directory @{text "FoobarCommand"}
 | 
| 66 | 948 | and add the line | 
| 949 | ||
| 207 | 950 |   @{text [display] "no_document use_thy \"Command\";"} 
 | 
| 66 | 951 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 952 |   to the file @{text "./FoobarCommand/ROOT.ML"}. You can now compile the theory by just typing:
 | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 953 | |
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 954 |   @{text [display] "$ isabelle make"}
 | 
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 955 | |
| 101 | 956 | If the compilation succeeds, you have finally created all the necessary log files. | 
| 957 | They are stored in the directory | |
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 958 | |
| 241 | 959 |   @{text [display]  "~/.isabelle/heaps/Isabelle2009/polyml-5.2.1_x86-linux/log"}
 | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 960 | |
| 74 | 961 | or something similar depending on your Isabelle distribution and architecture. | 
| 962 | One quick way to assign a shell variable to this directory is by typing | |
| 66 | 963 | |
| 964 |   @{text [display] "$ ISABELLE_LOGS=\"$(isabelle getenv -b ISABELLE_OUTPUT)\"/log"}
 | |
| 965 | ||
| 156 | 966 |   on the Unix prompt. If you now type @{text "ls $ISABELLE_LOGS"}, then the 
 | 
| 128 | 967 | directory should include the files: | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 968 | |
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 969 |   @{text [display] 
 | 
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 970 | "Pure.gz | 
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 971 | HOL.gz | 
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 972 | Pure-ProofGeneral.gz | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 973 | HOL-FoobarCommand.gz"} | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 974 | |
| 101 | 975 | From them you can create the keyword files. Assuming the name | 
| 75 | 976 |   of the directory is in  @{text "$ISABELLE_LOGS"},
 | 
| 74 | 977 | then the Unix command for creating the keyword file is: | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 978 | |
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 979 | @{text [display]
 | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 980 | "$ isabelle keywords -k foobar | 
| 80 | 981 |    $ISABELLE_LOGS/{Pure.gz,HOL.gz,Pure-ProofGeneral.gz,HOL-FoobarCommand.gz}"}
 | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 982 | |
| 80 | 983 |   The result is the file @{text "isar-keywords-foobar.el"}. It should contain
 | 
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 984 |   the string @{text "foobar"} twice.\footnote{To see whether things are fine,
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 985 |   check that @{text "grep foobar"} on this file returns something non-empty.}
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 986 |   This keyword file needs to be copied into the directory @{text
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 987 | "~/.isabelle/etc"}. To make ProofGeneral aware of it, you have to start | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 988 |   Isabelle with the option @{text "-k foobar"}, that is:
 | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 989 | |
| 80 | 990 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 991 |   @{text [display] "$ isabelle emacs -k foobar a_theory_file"}
 | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 992 | |
| 101 | 993 |   If you now build a theory on top of @{text "Command.thy"}, 
 | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 994 |   then you can use the command \isacommand{foobar}. You can just write
 | 
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 995 | *} | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 996 | |
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 997 | foobar | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 998 | |
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 999 | text {* 
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1000 | but you will not see any action as we chose to implement this command to do | 
| 327 
ce754ad78bc9
more work on the storing section
 Christian Urban <urbanc@in.tum.de> parents: 
326diff
changeset | 1001 | nothing. The point of this command is only to show the procedure of how | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 1002 | to interact with ProofGeneral. A similar procedure has to be done with any | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 1003 | other new command, and also any new keyword that is introduced with | 
| 327 
ce754ad78bc9
more work on the storing section
 Christian Urban <urbanc@in.tum.de> parents: 
326diff
changeset | 1004 |   the function @{ML_ind keyword in OuterKeyword}. For example:
 | 
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 1005 | *} | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1006 | |
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 1007 | ML{*val _ = OuterKeyword.keyword "blink" *}
 | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1008 | |
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 1009 | text {*
 | 
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1010 |   At the moment the command \isacommand{foobar} is not very useful. Let us
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1011 | refine it a bit next by letting it take a proposition as argument and | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1012 | printing this proposition inside the tracing buffer. | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1013 | |
| 75 | 1014 | The crucial part of a command is the function that determines the behaviour | 
| 1015 | of the command. In the code above we used a ``do-nothing''-function, which | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 1016 |   because of @{ML_ind succeed in Scan} does not parse any argument, but immediately
 | 
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 1017 |   returns the simple function @{ML "LocalTheory.theory I"}. We can
 | 
| 75 | 1018 | replace this code by a function that first parses a proposition (using the | 
| 1019 |   parser @{ML OuterParse.prop}), then prints out the tracing
 | |
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 1020 |   information (using a new function @{text trace_prop}) and 
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 1021 | finally does nothing. For this you can write: | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1022 | *} | 
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1023 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 1024 | ML{*let
 | 
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 1025 | fun trace_prop str = | 
| 327 
ce754ad78bc9
more work on the storing section
 Christian Urban <urbanc@in.tum.de> parents: 
326diff
changeset | 1026 | LocalTheory.theory (fn ctxt => (tracing str; ctxt)) | 
| 75 | 1027 | |
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1028 | val kind = OuterKeyword.thy_decl | 
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1029 | in | 
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1030 | OuterSyntax.local_theory "foobar_trace" "traces a proposition" | 
| 327 
ce754ad78bc9
more work on the storing section
 Christian Urban <urbanc@in.tum.de> parents: 
326diff
changeset | 1031 | kind (OuterParse.prop >> trace_prop) | 
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 1032 | end *} | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1033 | |
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1034 | text {*
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1035 |   The command is now \isacommand{foobar\_trace} and can be used to 
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1036 | see the proposition in the tracing buffer. | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1037 | *} | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1038 | |
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1039 | foobar_trace "True \<and> False" | 
| 218 
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
 Christian Urban <urbanc@in.tum.de> parents: 
211diff
changeset | 1040 | |
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1041 | text {*
 | 
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 1042 |   Note that so far we used @{ML_ind thy_decl in OuterKeyword} as the kind
 | 
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 1043 | indicator for the command. This means that the command finishes as soon as | 
| 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 1044 | the arguments are processed. Examples of this kind of commands are | 
| 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 1045 |   \isacommand{definition} and \isacommand{declare}.  In other cases, commands
 | 
| 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 1046 | are expected to parse some arguments, for example a proposition, and then | 
| 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 1047 | ``open up'' a proof in order to prove the proposition (for example | 
| 86 | 1048 |   \isacommand{lemma}) or prove some other properties (for example
 | 
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 1049 |   \isacommand{function}). To achieve this kind of behaviour, you have to use
 | 
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 1050 |   the kind indicator @{ML_ind thy_goal in OuterKeyword} and the function @{ML
 | 
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 1051 | "local_theory_to_proof" in OuterSyntax} to set up the command. Note, | 
| 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 1052 |   however, once you change the ``kind'' of a command from @{ML thy_decl in
 | 
| 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 1053 |   OuterKeyword} to @{ML thy_goal in OuterKeyword} then the keyword file needs
 | 
| 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 1054 | to be re-created! | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1055 | |
| 327 
ce754ad78bc9
more work on the storing section
 Christian Urban <urbanc@in.tum.de> parents: 
326diff
changeset | 1056 |   Below we show the command \isacommand{foobar\_goal} which takes a
 | 
| 
ce754ad78bc9
more work on the storing section
 Christian Urban <urbanc@in.tum.de> parents: 
326diff
changeset | 1057 | proposition as argument and then starts a proof in order to prove | 
| 
ce754ad78bc9
more work on the storing section
 Christian Urban <urbanc@in.tum.de> parents: 
326diff
changeset | 1058 |   it. Therefore in Line 9, we set the kind indicator to @{ML thy_goal in
 | 
| 
ce754ad78bc9
more work on the storing section
 Christian Urban <urbanc@in.tum.de> parents: 
326diff
changeset | 1059 | OuterKeyword}. | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1060 | *} | 
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1061 | |
| 114 
13fd0a83d3c3
properly handled linenumbers in ML-text and Isar-proofs
 Christian Urban <urbanc@in.tum.de> parents: 
108diff
changeset | 1062 | ML%linenosgray{*let
 | 
| 327 
ce754ad78bc9
more work on the storing section
 Christian Urban <urbanc@in.tum.de> parents: 
326diff
changeset | 1063 | fun goal_prop str lthy = | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1064 | let | 
| 241 | 1065 | val prop = Syntax.read_prop lthy str | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1066 | in | 
| 241 | 1067 | Proof.theorem_i NONE (K I) [[(prop,[])]] lthy | 
| 327 
ce754ad78bc9
more work on the storing section
 Christian Urban <urbanc@in.tum.de> parents: 
326diff
changeset | 1068 | end | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1069 | |
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1070 | val kind = OuterKeyword.thy_goal | 
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1071 | in | 
| 327 
ce754ad78bc9
more work on the storing section
 Christian Urban <urbanc@in.tum.de> parents: 
326diff
changeset | 1072 | OuterSyntax.local_theory_to_proof "foobar_goal" "proves a proposition" | 
| 
ce754ad78bc9
more work on the storing section
 Christian Urban <urbanc@in.tum.de> parents: 
326diff
changeset | 1073 | kind (OuterParse.prop >> goal_prop) | 
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 1074 | end *} | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1075 | |
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1076 | text {*
 | 
| 327 
ce754ad78bc9
more work on the storing section
 Christian Urban <urbanc@in.tum.de> parents: 
326diff
changeset | 1077 |   The function @{text goal_prop} in Lines 2 to 7 takes a string (the proposition to be
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 1078 | proved) and a context as argument. The context is necessary in order to be able to use | 
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 1079 |   @{ML_ind read_prop in Syntax}, which converts a string into a proper proposition.
 | 
| 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 1080 |   In Line 6 the function @{ML_ind theorem_i in Proof} starts the proof for the
 | 
| 75 | 1081 |   proposition. Its argument @{ML NONE} stands for a locale (which we chose to
 | 
| 1082 |   omit); the argument @{ML "(K I)"} stands for a function that determines what
 | |
| 1083 | should be done with the theorem once it is proved (we chose to just forget | |
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 1084 | about it). Line 9 contains the parser for the proposition. | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 1085 | |
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1086 |   If you now type \isacommand{foobar\_goal}~@{text [quotes] "True \<and> True"},
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1087 | you obtain the following proof state | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1088 | *} | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1089 | |
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1090 | foobar_goal "True \<and> True" | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1091 | txt {*
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1092 |   \begin{minipage}{\textwidth}
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1093 |   @{subgoals [display]}
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1094 |   \end{minipage}\medskip
 | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1095 | |
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1096 | and can prove the proposition as follows. | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1097 | *} | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1098 | apply(rule conjI) | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1099 | apply(rule TrueI)+ | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1100 | done | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1101 | |
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1102 | text {*
 | 
| 327 
ce754ad78bc9
more work on the storing section
 Christian Urban <urbanc@in.tum.de> parents: 
326diff
changeset | 1103 |   {\bf TBD below}
 | 
| 74 | 1104 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 1105 |   (FIXME: read a name and show how to store theorems; see @{ML_ind note in LocalTheory})
 | 
| 241 | 1106 | |
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1107 | *} | 
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1108 | |
| 328 
c0cae24b9d46
updated to new Isabelle; more work on the data section
 Christian Urban <urbanc@in.tum.de> parents: 
327diff
changeset | 1109 | ML_val{*val r = Unsynchronized.ref (NONE:(unit -> term) option)*}
 | 
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1110 | ML{*let
 | 
| 324 
4172c0743cf2
updated foobar_proof example
 Christian Urban <urbanc@in.tum.de> parents: 
322diff
changeset | 1111 | fun after_qed thm_name thms lthy = | 
| 
4172c0743cf2
updated foobar_proof example
 Christian Urban <urbanc@in.tum.de> parents: 
322diff
changeset | 1112 | LocalTheory.note Thm.theoremK (thm_name, (flat thms)) lthy |> snd | 
| 
4172c0743cf2
updated foobar_proof example
 Christian Urban <urbanc@in.tum.de> parents: 
322diff
changeset | 1113 | |
| 
4172c0743cf2
updated foobar_proof example
 Christian Urban <urbanc@in.tum.de> parents: 
322diff
changeset | 1114 | fun setup_proof (thm_name, (txt, pos)) lthy = | 
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1115 | let | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1116 |      val trm = ML_Context.evaluate lthy true ("r", r) txt
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1117 | in | 
| 324 
4172c0743cf2
updated foobar_proof example
 Christian Urban <urbanc@in.tum.de> parents: 
322diff
changeset | 1118 | Proof.theorem_i NONE (after_qed thm_name) [[(trm,[])]] lthy | 
| 
4172c0743cf2
updated foobar_proof example
 Christian Urban <urbanc@in.tum.de> parents: 
322diff
changeset | 1119 | end | 
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1120 | |
| 324 
4172c0743cf2
updated foobar_proof example
 Christian Urban <urbanc@in.tum.de> parents: 
322diff
changeset | 1121 | val parser = SpecParse.opt_thm_name ":" -- OuterParse.ML_source | 
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1122 | in | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1123 | OuterSyntax.local_theory_to_proof "foobar_prove" "proving a proposition" | 
| 324 
4172c0743cf2
updated foobar_proof example
 Christian Urban <urbanc@in.tum.de> parents: 
322diff
changeset | 1124 | OuterKeyword.thy_goal (parser >> setup_proof) | 
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1125 | end*} | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1126 | |
| 324 
4172c0743cf2
updated foobar_proof example
 Christian Urban <urbanc@in.tum.de> parents: 
322diff
changeset | 1127 | foobar_prove test: {* @{prop "True"} *}
 | 
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1128 | apply(rule TrueI) | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1129 | done | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1130 | |
| 322 | 1131 | (* | 
| 1132 | ML {*
 | |
| 1133 | structure TacticData = ProofDataFun | |
| 1134 | ( | |
| 1135 | type T = thm list -> tactic; | |
| 1136 | fun init _ = undefined; | |
| 1137 | ); | |
| 1138 | ||
| 1139 | val set_tactic = TacticData.put; | |
| 1140 | *} | |
| 1141 | ||
| 1142 | ML {*
 | |
| 1143 |   TacticData.get @{context}
 | |
| 1144 | *} | |
| 1145 | ||
| 1146 | ML {* Method.set_tactic  *}
 | |
| 1147 | ML {* fun tactic (facts: thm list) : tactic = (atac 1) *}
 | |
| 1148 | ML {* Context.map_proof *}
 | |
| 1149 | ML {* ML_Context.expression *}
 | |
| 1150 | ML {* METHOD *}
 | |
| 1151 | ||
| 1152 | ||
| 1153 | ML {* 
 | |
| 1154 | fun myexpression pos bind body txt = | |
| 1155 | let | |
| 1156 |   val _ = tracing ("bind)" ^ bind)
 | |
| 1157 |   val _ = tracing ("body)" ^ body)
 | |
| 1158 |   val _ = tracing ("txt)"  ^ txt)
 | |
| 1159 |   val _ = tracing ("result) " ^ "Context.set_thread_data (SOME (let " ^ bind ^ " = " ^ txt ^ " in " ^ body ^
 | |
| 1160 | " end (ML_Context.the_generic_context ())));") | |
| 1161 | in | |
| 1162 | ML_Context.exec (fn () => ML_Context.eval false pos | |
| 1163 |     ("Context.set_thread_data (SOME (let " ^ bind ^ " = " ^ txt ^ " in " ^ body ^
 | |
| 1164 | " end (ML_Context.the_generic_context ())));")) | |
| 1165 | end | |
| 1166 | *} | |
| 319 
6bce4acf7f2a
added file for producing a keyword file
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 1167 | |
| 
6bce4acf7f2a
added file for producing a keyword file
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 1168 | |
| 322 | 1169 | ML {*
 | 
| 1170 | fun ml_tactic (txt, pos) ctxt = | |
| 1171 | let | |
| 1172 | val ctxt' = ctxt |> Context.proof_map | |
| 1173 | (myexpression pos | |
| 1174 | "fun tactic (facts: thm list) : tactic" | |
| 1175 | "Context.map_proof (Method.set_tactic tactic)" txt); | |
| 1176 | in | |
| 1177 | Context.setmp_thread_data (SOME (Context.Proof ctxt)) (TacticData.get ctxt') | |
| 1178 | end; | |
| 1179 | *} | |
| 1180 | ||
| 1181 | ML {*
 | |
| 1182 | fun tactic3 (txt, pos) ctxt = | |
| 1183 | let | |
| 1184 |     val _ = tracing ("1) " ^ txt )
 | |
| 1185 | in | |
| 1186 | METHOD (ml_tactic (txt, pos) ctxt; K (atac 1)) | |
| 1187 | end | |
| 1188 | *} | |
| 1189 | ||
| 1190 | setup {*
 | |
| 1191 | Method.setup (Binding.name "tactic3") (Scan.lift (OuterParse.position Args.name) | |
| 1192 | >> tactic3) | |
| 1193 | "ML tactic as proof method" | |
| 1194 | *} | |
| 1195 | ||
| 1196 | lemma "A \<Longrightarrow> A" | |
| 1197 | apply(tactic3 {* (atac 1)  *})
 | |
| 1198 | done | |
| 1199 | ||
| 1200 | ML {* 
 | |
| 1201 | (ML_Context.the_generic_context ()) | |
| 1202 | *} | |
| 1203 | ||
| 1204 | ML {*
 | |
| 1205 | Context.set_thread_data; | |
| 1206 | ML_Context.the_generic_context | |
| 1207 | *} | |
| 1208 | ||
| 1209 | lemma "A \<Longrightarrow> A" | |
| 1210 | ML_prf {*
 | |
| 1211 | Context.set_thread_data (SOME (let fun tactic (facts: thm list) : tactic = (atac 1) in Context.map_proof (Method.set_tactic tactic) end (ML_Context.the_generic_context ()))); | |
| 1212 | *} | |
| 1213 | ||
| 1214 | ML {*
 | |
| 1215 | Context.set_thread_data (SOME ((let fun tactic (facts: thm list) : tactic = (atac 1) in 3 end) (ML_Context.the_generic_context ()))); | |
| 1216 | *} | |
| 1217 | ||
| 1218 | ML {*
 | |
| 1219 | Context.set_thread_data (SOME (let | |
| 1220 | fun tactic (facts: thm list) : tactic = (atac 1) | |
| 1221 | in | |
| 1222 | Context.map_proof (Method.set_tactic tactic) | |
| 1223 | end | |
| 1224 | (ML_Context.the_generic_context ()))); | |
| 1225 | *} | |
| 1226 | ||
| 1227 | ||
| 1228 | ML {*
 | |
| 1229 | let | |
| 1230 | fun tactic (facts: thm list) : tactic = atac | |
| 1231 | in | |
| 1232 | Context.map_proof (Method.set_tactic tactic) | |
| 1233 | end *} | |
| 1234 | ||
| 1235 | end *} | |
| 1236 | ||
| 1237 | ML {* Toplevel.program (fn () => 
 | |
| 1238 | (ML_Context.expression Position.none "val plus : int" "3 + 4" "1" (Context.Proof @{context})))*}
 | |
| 1239 | ||
| 1240 | ||
| 1241 | ML {*
 | |
| 1242 | fun ml_tactic (txt, pos) ctxt = | |
| 1243 | let | |
| 1244 | val ctxt' = ctxt |> Context.proof_map | |
| 1245 | (ML_Context.expression pos | |
| 1246 | "fun tactic (facts: thm list) : tactic" | |
| 1247 | "Context.map_proof (Method.set_tactic tactic)" txt); | |
| 1248 | in Context.setmp_thread_data (SOME (Context.Proof ctxt)) (TacticData.get ctxt') end; | |
| 1249 | ||
| 1250 | *} | |
| 1251 | ||
| 1252 | ML {*
 | |
| 1253 | Context.set_thread_data (SOME (let fun tactic (facts: thm list) : tactic = (atac 1) in Context.map_proof (Method.set_tactic tactic) end (ML_Context.the_generic_context ()))); | |
| 1254 | *} | |
| 1255 | *) | |
| 319 
6bce4acf7f2a
added file for producing a keyword file
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 1256 | |
| 211 
d5accbc67e1b
more work on simple inductive and marked all sections that are still seriously incomplete with TBD
 Christian Urban <urbanc@in.tum.de> parents: 
207diff
changeset | 1257 | section {* Methods (TBD) *}
 | 
| 178 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1258 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1259 | text {*
 | 
| 207 | 1260 | (FIXME: maybe move to after the tactic section) | 
| 1261 | ||
| 221 | 1262 | Methods are central to Isabelle. They are the ones you use for example | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
183diff
changeset | 1263 |   in \isacommand{apply}. To print out all currently known methods you can use the 
 | 
| 192 | 1264 | Isabelle command: | 
| 178 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1265 | |
| 207 | 1266 |   \begin{isabelle}
 | 
| 1267 |   \isacommand{print\_methods}\\
 | |
| 1268 |   @{text "> methods:"}\\
 | |
| 1269 |   @{text ">   -:  do nothing (insert current facts only)"}\\
 | |
| 1270 |   @{text ">   HOL.default:  apply some intro/elim rule (potentially classical)"}\\
 | |
| 1271 |   @{text ">   ..."}
 | |
| 1272 |   \end{isabelle}
 | |
| 178 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1273 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1274 | An example of a very simple method is: | 
| 178 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1275 | *} | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1276 | |
| 244 
dc95a56b1953
fixed the problem with double definition of even and odd
 Christian Urban <urbanc@in.tum.de> parents: 
241diff
changeset | 1277 | method_setup %gray foo = | 
| 181 
5baaabe1ab92
updated to new method_setup
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 1278 |  {* Scan.succeed
 | 
| 
5baaabe1ab92
updated to new method_setup
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 1279 |       (K (SIMPLE_METHOD ((etac @{thm conjE} THEN' rtac @{thm conjI}) 1))) *}
 | 
| 244 
dc95a56b1953
fixed the problem with double definition of even and odd
 Christian Urban <urbanc@in.tum.de> parents: 
241diff
changeset | 1280 | "foo method for conjE and conjI" | 
| 178 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1281 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1282 | text {*
 | 
| 286 | 1283 |   It defines the method @{text foo}, which takes no arguments (therefore the
 | 
| 207 | 1284 |   parser @{ML Scan.succeed}) and only applies a single tactic, namely the tactic which 
 | 
| 256 
1fb8d62c88a0
added some first index-information
 Christian Urban <urbanc@in.tum.de> parents: 
250diff
changeset | 1285 |   applies @{thm [source] conjE} and then @{thm [source] conjI}. The function 
 | 
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 1286 |   @{ML_ind SIMPLE_METHOD in Method}
 | 
| 287 | 1287 |   turns such a tactic into a method. The method @{text "foo"} can be used as follows
 | 
| 178 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1288 | *} | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1289 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1290 | lemma shows "A \<and> B \<Longrightarrow> C \<and> D" | 
| 244 
dc95a56b1953
fixed the problem with double definition of even and odd
 Christian Urban <urbanc@in.tum.de> parents: 
241diff
changeset | 1291 | apply(foo) | 
| 178 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1292 | txt {*
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1293 | where it results in the goal state | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1294 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1295 |   \begin{minipage}{\textwidth}
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1296 |   @{subgoals}
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1297 |   \end{minipage} *}
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1298 | (*<*)oops(*>*) | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1299 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1300 | |
| 319 
6bce4acf7f2a
added file for producing a keyword file
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 1301 | |
| 
6bce4acf7f2a
added file for producing a keyword file
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 1302 | |
| 
6bce4acf7f2a
added file for producing a keyword file
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 1303 | |
| 
6bce4acf7f2a
added file for producing a keyword file
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 1304 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1305 | (* | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1306 | ML {* SIMPLE_METHOD *}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1307 | ML {* METHOD *}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1308 | ML {* K (SIMPLE_METHOD ((etac @{thm conjE} THEN' rtac @{thm conjI}) 1)) *}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1309 | ML {* Scan.succeed  *}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1310 | *) | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1311 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
183diff
changeset | 1312 | text {*
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
183diff
changeset | 1313 | (FIXME: explain a version of rule-tac) | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
183diff
changeset | 1314 | *} | 
| 178 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1315 | |
| 75 | 1316 | (*<*) | 
| 194 | 1317 | (* THIS IS AN OLD VERSION OF THE PARSING CHAPTER BY JEREMY DAWSON *) | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 1318 | chapter {* Parsing *}
 | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 1319 | |
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 1320 | text {*
 | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 1321 | |
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1322 | 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 | 1323 | including: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1324 |   \begin{itemize}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1325 | \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 | 1326 | or simplified versions of such code | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1327 | \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 | 1328 | (or specialisations of their types) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1329 | \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 | 1330 | 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 | 1331 | \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 | 1332 | 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 | 1333 | \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 | 1334 |   \end{itemize}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1335 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1336 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1337 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1338 | section {* Parsing Isar input *}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1339 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1340 | text {*
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1341 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1342 | The typical parsing function has the type | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1343 |   \texttt{'src -> 'res * 'src}, with input  
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1344 |   of type \texttt{'src}, returning a result 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1345 |   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 | 1346 | 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 | 1347 | (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 | 1348 | 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 | 1349 |   value of type \texttt{'res}). 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1350 | An exception is raised if an appropriate value | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1351 | cannot be produced from the input. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1352 | 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 | 1353 | for the failure of a parse. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1354 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1355 | 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 | 1356 | which is of type | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1357 |   \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 | 1358 |   (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 | 1359 | However, much of the discussion at | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1360 | 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 | 1361 | is relevant. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1362 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1363 | 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 | 1364 | as follows: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1365 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1366 | open StringCvt ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1367 |   type ('res, 'src) ex_reader = 'src -> 'res * 'src
 | 
| 75 | 1368 |   ex_reader : ('res, 'src) reader -> ('res, 'src) ex_reader 
 | 
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1369 | fun ex_reader rdr src = Option.valOf (rdr src) ; | 
| 75 | 1370 |   reader : ('res, 'src) ex_reader -> ('res, 'src) reader 
 | 
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1371 | 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 | 1372 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1373 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1374 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1375 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1376 | section{* The \texttt{Scan} structure *}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1377 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1378 | text {* 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1379 |   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 | 1380 | 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 | 1381 |   of the type \texttt{'src -> 'res * 'src}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1382 | Three exceptions are used: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1383 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1384 | 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 | 1385 | 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 | 1386 | exception ABORT of string; (*dead end*) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1387 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1388 | 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 | 1389 | symbols) are declared as infix. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1390 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1391 | Some functions from that structure are | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1392 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1393 |   |-- : ('src -> 'res1 * 'src') * ('src' -> 'res2 * 'src'') ->
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1394 | 'src -> 'res2 * 'src'' | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1395 |   --| : ('src -> 'res1 * 'src') * ('src' -> 'res2 * 'src'') ->
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1396 | 'src -> 'res1 * 'src'' | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1397 |   -- : ('src -> 'res1 * 'src') * ('src' -> 'res2 * 'src'') ->
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1398 |   'src -> ('res1 * 'res2) * 'src''
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1399 |   ^^ : ('src -> string * 'src') * ('src' -> string * 'src'') ->
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1400 | 'src -> string * 'src'' | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1401 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1402 | 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 | 1403 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1404 |   \texttt{|--} and \texttt{--|} 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1405 | 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 | 1406 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1407 |   \texttt{--} returns both.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1408 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1409 | \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 | 1410 | (which must be strings). | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1411 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1412 | Note how, although the types | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1413 |   \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 | 1414 | 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 | 1415 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1416 |   :-- : ('src -> 'res1 * 'src') * ('res1 -> 'src' -> 'res2 * 'src'') ->
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1417 |   'src -> ('res1 * 'res2) * 'src''
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1418 |   :|-- : ('src -> 'res1 * 'src') * ('res1 -> 'src' -> 'res2 * 'src'') ->
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1419 | 'src -> 'res2 * 'src'' | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1420 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1421 |   These are similar to \texttt{|--} and \texttt{--|},
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1422 | 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 | 1423 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1424 |   >> : ('src -> 'res1 * 'src') * ('res1 -> 'res2) -> 'src -> 'res2 * 'src'
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1425 |   || : ('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 | 1426 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1427 |   \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 | 1428 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1429 |   \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 | 1430 |   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 | 1431 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1432 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1433 |   succeed : 'res -> ('src -> 'res * 'src) ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1434 |   fail : ('src -> 'res_src) ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1435 |   !! : ('src * string option -> string) -> 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1436 |   ('src -> 'res_src) -> ('src -> 'res_src) ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1437 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1438 |   \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 | 1439 |   \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 | 1440 |   \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 | 1441 |   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 | 1442 | 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 | 1443 |   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 | 1444 |   it won't recover by trying \texttt{parse2}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1445 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1446 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1447 |   one : ('si -> bool) -> ('si list -> 'si * 'si list) ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1448 |   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 | 1449 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1450 | 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 | 1451 |   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 | 1452 |   On other failures they raise \texttt{FAIL NONE} 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1453 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1454 |   \texttt{one p} takes the first
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1455 |   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 | 1456 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1457 |   \texttt{some f} takes the first
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1458 |   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 | 1459 |   \texttt{NONE}.  
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1460 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1461 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1462 |   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 | 1463 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1464 |   \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 | 1465 |   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 | 1466 |   it fails, raising \texttt{MORE NONE}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1467 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1468 |   \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 | 1469 |   does not satisfy \texttt{p}.  
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1470 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1471 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1472 |   option : ('src -> 'res * 'src) -> ('src -> 'res option * 'src)
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1473 |   optional : ('src -> 'res * 'src) -> 'res -> ('src -> 'res * 'src)
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1474 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1475 |   \texttt{option}: 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1476 |   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 | 1477 |   or raises \texttt{FAIL \_},
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1478 |   \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 | 1479 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1480 |   \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 | 1481 |   \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 | 1482 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1483 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1484 |   repeat : ('src -> 'res * 'src) -> 'src -> 'res list * 'src
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1485 |   repeat1 : ('src -> 'res * 'src) -> 'src -> 'res list * 'src
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1486 |   bulk : ('src -> 'res * 'src) -> 'src -> 'res list * 'src 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1487 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1488 |   \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 | 1489 |   \texttt{f} fails with \texttt{FAIL \_}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1490 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1491 |   \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 | 1492 | successful parse. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1493 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1494 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1495 |   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 | 1496 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1497 |   \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 | 1498 |   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 | 1499 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1500 |   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 | 1501 | HOW DO THEY WORK ?? TO BE COMPLETED | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1502 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1503 | dest_lexicon: lexicon -> string list ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1504 | make_lexicon: string list list -> lexicon ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1505 | empty_lexicon: lexicon ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1506 | extend_lexicon: string list list -> lexicon -> lexicon ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1507 | merge_lexicons: lexicon -> lexicon -> lexicon ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1508 | is_literal: lexicon -> string list -> bool ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1509 | literal: lexicon -> string list -> string list * string list ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1510 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1511 | 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 | 1512 | by: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1513 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1514 | val (command_lexicon, keyword_lexicon) = OuterSyntax.get_lexicons () ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1515 | val commands = Scan.dest_lexicon command_lexicon ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1516 | val keywords = Scan.dest_lexicon keyword_lexicon ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1517 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1518 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1519 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1520 | section{* The \texttt{OuterLex} structure *}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1521 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1522 | text {*
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1523 |   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 | 1524 | 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 | 1525 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1526 | structure T = OuterLex; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1527 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1528 |   This structure defines the type \texttt{token}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1529 | (The types | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1530 |   \texttt{OuterLex.token},
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1531 |   \texttt{OuterParse.token} and
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1532 |   \texttt{SpecParse.token} are all the same).
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1533 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1534 | 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 | 1535 |   functions is \texttt{token list}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1536 | |
| 250 
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
 Christian Urban <urbanc@in.tum.de> parents: 
244diff
changeset | 1537 | The datatype definition (which is not published in the signature) is | 
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1538 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1539 | 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 | 1540 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1541 | 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 | 1542 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1543 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1544 | |
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 1545 | |
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 1546 | |
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1547 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 1548 | ML{*
 | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 1549 | 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 | 1550 |    "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 | 1551 | *} | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 1552 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 1553 | ML{*
 | 
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1554 | 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 | 1555 | *} | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 1556 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 1557 | ML{*
 | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 1558 | 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 | 1559 | *} | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 1560 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 1561 | ML{*
 | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 1562 | 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 | 1563 | *} | 
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1564 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 1565 | ML{*
 | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 1566 | 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 | 1567 | *} | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 1568 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 1569 | ML{*
 | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 1570 | 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 | 1571 | *} | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 1572 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 1573 | ML{*
 | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 1574 | OuterLex.stopper | 
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1575 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1576 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1577 | text {*
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1578 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1579 |   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 | 1580 | 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 | 1581 | comments to have been filtered out. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1582 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1583 | There is a special end-of-file token: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1584 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1585 | 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 | 1586 | (* end of file token *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1587 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1588 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1589 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1590 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1591 | section {* The \texttt{OuterParse} structure *}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1592 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1593 | text {*
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1594 |   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 | 1595 | 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 | 1596 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1597 | structure P = OuterParse; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1598 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1599 |   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 | 1600 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1601 | 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 | 1602 |   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 | 1603 |   \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 | 1604 | 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 | 1605 | 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 | 1606 |   (eg, \texttt{name, xname, text, typ}).
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1607 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1608 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1609 | 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 | 1610 | $$$ : string -> string tlp | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1611 | nat : int tlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1612 | maybe : 'a tlp -> 'a option tlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1613 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1614 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1615 |   \texttt{\$\$\$ s} returns the first token,
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1616 |   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 | 1617 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1618 |   \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 | 1619 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1620 |   \texttt{maybe}: if \texttt{p} returns \texttt{r}, 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1621 |   then \texttt{maybe p} returns \texttt{SOME r} ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1622 |   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 | 1623 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1624 | A few examples: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1625 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1626 | 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 | 1627 | 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 | 1628 | 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 | 1629 | 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 | 1630 | 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 | 1631 | 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 | 1632 | 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 | 1633 | 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 | 1634 | 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 | 1635 | 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 | 1636 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1637 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1638 | The following code helps run examples: | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1639 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1640 | fun parse_str tlp str = | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1641 | let val toks : token list = OuterSyntax.scan str ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1642 | 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 | 1643 | val (res, rem_toks) = tlp proper_toks ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1644 | val rem_str = String.concat | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1645 | (Library.separate " " (List.map T.unparse rem_toks)) ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1646 | in (res, rem_str) end ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1647 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1648 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1649 |   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 | 1650 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1651 | val type_args = | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1652 | type_ident >> Library.single || | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1653 |   $$$ "(" |-- !!! (list1 type_ident --| $$$ ")") ||
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1654 | Scan.succeed []; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1655 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1656 | 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 | 1657 | 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 | 1658 | list. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1659 |   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 | 1660 | 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 | 1661 | and then a ")" which is also ignored. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1662 |   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 | 1663 |   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 | 1664 | 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 | 1665 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1666 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1667 | fun triple2 (x, (y, z)) = (x, y, z); | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1668 | val arity = xname -- ($$$ "::" |-- !!! ( | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1669 |   Scan.optional ($$$ "(" |-- !!! (list1 sort --| $$$ ")")) []
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1670 | -- sort)) >> triple2; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1671 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1672 |   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 | 1673 | 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 | 1674 |   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 | 1675 | 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 | 1676 |   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 | 1677 | 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 | 1678 |   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 | 1679 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1680 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1681 |   parse_str P.type_args "('a, 'b) ntyp" ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1682 | parse_str P.type_args "'a ntyp" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1683 | parse_str P.type_args "ntyp" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1684 | parse_str P.arity "ty :: tycl" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1685 | parse_str P.arity "ty :: (tycl1, tycl2) tycl" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1686 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1687 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1688 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1689 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1690 | section {* The \texttt{SpecParse} structure *}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1691 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1692 | text {*
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1693 |   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 | 1694 | 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 | 1695 | For example, | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1696 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1697 | open SpecParse ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1698 | attrib : Attrib.src tok_rdr ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1699 | attribs : Attrib.src list tok_rdr ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1700 | opt_attribs : Attrib.src list tok_rdr ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1701 | xthm : (thmref * Attrib.src list) tok_rdr ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1702 | xthms1 : (thmref * Attrib.src list) list tok_rdr ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1703 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1704 | parse_str attrib "simp" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1705 | parse_str opt_attribs "hello" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1706 | 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 | 1707 | map Args.dest_src ass ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1708 | 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 | 1709 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1710 | parse_str xthm "mythm [attr]" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1711 | parse_str xthms1 "thm1 [attr] thms2" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1712 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1713 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1714 |   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 | 1715 | structure, described below. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1716 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1717 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1718 | section{* The \texttt{Args} structure *}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1719 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1720 | text {*
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1721 |   The source file is \texttt{src/Pure/Isar/args.ML}.
 | 
| 250 
ab9e09076462
some polishing; added together with Jasmin more examples to the pretty printing section
 Christian Urban <urbanc@in.tum.de> parents: 
244diff
changeset | 1722 |   The primary type of this structure is the \texttt{src} datatype;
 | 
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1723 | 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 | 1724 |   \texttt{Args.src} and \texttt{Args.dest\_src}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1725 | are in fact the constructor and destructor functions. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1726 |   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 | 1727 |   are in fact \texttt{Args.src}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1728 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1729 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1730 | 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 | 1731 | 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 | 1732 | 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 | 1733 | 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 | 1734 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1735 | val thy = ML_Context.the_context () ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1736 | val ctxt = ProofContext.init thy ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1737 | map (pr_src ctxt) ass ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1738 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1739 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1740 |   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 | 1741 |   ``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 | 1742 | input. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1743 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1744 | (* how an Args.src is parsed *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1745 |   P.position : 'a tlp -> ('a * Position.T) tlp ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1746 | P.arguments : Args.T list tlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1747 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1748 | val parse_src : Args.src tlp = | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1749 | P.position (P.xname -- P.arguments) >> Args.src ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1750 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1751 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1752 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1753 | 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 | 1754 | map Args.string_of args ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1755 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1756 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1757 |   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 | 1758 |   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 | 1759 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1760 | 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 | 1761 | open Args ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1762 | nat : int atlp ; (* also Args.int *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1763 | thm_sel : PureThy.interval list atlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1764 | list : 'a atlp -> 'a list atlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1765 | attribs : (string -> string) -> Args.src list atlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1766 | opt_attribs : (string -> string) -> Args.src list atlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1767 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1768 | (* parse_atl_str : 'a atlp -> (string -> 'a * string) ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1769 | 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 | 1770 | fun parse_atl_str atlp str = | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1771 | 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 | 1772 | val (res, rem_ats) = atlp ats ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1773 | in (res, String.concat (Library.separate " " | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1774 | (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 | 1775 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1776 | parse_atl_str Args.int "-1-," ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1777 | parse_atl_str (Scan.option Args.int) "x1-," ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1778 | 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 | 1779 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1780 | 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 | 1781 | "[THEN trans [THEN sym], simp, OF sym]" ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1782 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1783 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1784 |   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 | 1785 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1786 |   \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 | 1787 | and also refer to a generic context. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1788 |   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 | 1789 |   (as does \texttt{Attrib} - RETHINK THIS)
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1790 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1791 |   (\texttt{Args.syntax} shown below has type specialised)
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1792 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1793 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1794 |   type ('res, 'src) parse_fn = 'src -> 'res * 'src ;
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1795 |   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 | 1796 | Scan.lift : 'a atlp -> 'a cgatlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1797 | term : term cgatlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1798 | typ : typ cgatlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1799 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1800 |   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 | 1801 | Attrib.thm : thm cgatlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1802 | Attrib.thms : thm list cgatlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1803 | Attrib.multi_thm : thm list cgatlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1804 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1805 | (* parse_cgatl_str : 'a cgatlp -> (string -> 'a * string) ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1806 | 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 | 1807 | fun parse_cgatl_str cgatlp str = | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1808 | let | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1809 | (* use the current generic context *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1810 | val generic = Context.Theory thy ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1811 | 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 | 1812 | (* ignore any change to the generic context *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1813 | val (res, (_, rem_ats)) = cgatlp (generic, ats) ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1814 | in (res, String.concat (Library.separate " " | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1815 | (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 | 1816 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1817 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1818 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1819 | section{* Attributes, and the \texttt{Attrib} structure *}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1820 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1821 | text {*
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1822 |   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 | 1823 |   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 | 1824 |   \texttt{src/Pure/Isar/attrib.ML}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1825 | 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 | 1826 | 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 | 1827 | 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 | 1828 | theory). | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1829 |   The functions \texttt{Thm.rule\_attribute} and
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1830 |   \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 | 1831 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1832 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1833 | type attribute = Context.generic * thm -> Context.generic * thm; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1834 | 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 | 1835 | Thm.rule_attribute : (Context.generic -> thm -> thm) -> attribute ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1836 | Thm.declaration_attribute : (thm -> Context.generic trf) -> attribute ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1837 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1838 | Attrib.print_attributes : theory -> unit ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1839 | 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 | 1840 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1841 | 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 | 1842 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1843 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1844 | 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 | 1845 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1846 | Attrib.add_attributes : | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1847 | (bstring * (src -> attribute) * string) list -> theory trf ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1848 | (* | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1849 |   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 | 1850 | *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1851 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1852 | 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 | 1853 | 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 | 1854 | (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 | 1855 |   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 | 1856 |   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 | 1857 |   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 | 1858 |   \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 | 1859 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1860 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1861 | FullAttrib.THEN_att : src -> attribute ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1862 | 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 | 1863 | 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 | 1864 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1865 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1866 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1867 | Attrib.syntax : attribute cgatlp -> src -> attribute ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1868 | Attrib.no_args : attribute -> src -> attribute ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1869 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1870 |   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 | 1871 |   the generic context \texttt{gc} is used 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1872 |   (and potentially changed to \texttt{gc'})
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1873 |   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 | 1874 |   then be applied to \texttt{(gc', th)}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1875 |   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 | 1876 | which must all be consumed by the parse. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1877 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1878 |   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 | 1879 |   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 | 1880 |   \texttt{src} must be empty.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1881 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1882 |   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 | 1883 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1884 | 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 | 1885 | rot_att_n : int -> attribute ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1886 | 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 | 1887 | val rotated_att : src -> attribute = | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1888 | 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 | 1889 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1890 | val THEN_arg : int cgatlp = Scan.lift | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1891 | (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 | 1892 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1893 | Attrib.thm : thm cgatlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1894 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1895 | THEN_arg -- Attrib.thm : (int * thm) cgatlp ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1896 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1897 | 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 | 1898 | THEN_att_n : int * thm -> attribute ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1899 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1900 | val THEN_att : src -> attribute = Attrib.syntax | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1901 | (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 | 1902 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1903 |   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 | 1904 |   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 | 1905 |   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 | 1906 | 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 | 1907 |   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 | 1908 |   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 | 1909 |   parsed by \texttt{Attrib.thm}.  
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1910 |   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 | 1911 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1912 | *} | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1913 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1914 | section{* Methods, and the \texttt{Method} structure *}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1915 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1916 | text {*
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1917 |   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 | 1918 |   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 | 1919 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1920 | (* datatype method = Meth of thm list -> cases_tactic; *) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1921 | RuleCases.NO_CASES : tactic -> cases_tactic ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1922 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1923 |   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 | 1924 |   \texttt{Meth}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1925 |   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 | 1926 |   \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 | 1927 |   \texttt{cases\_tactic} without any further case information.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1928 |   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 | 1929 | 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 | 1930 |   \emph{facts} in the proof.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1931 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1932 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1933 | RAW_METHOD : (thm list -> tactic) -> method ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1934 | METHOD : (thm list -> tactic) -> method ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1935 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1936 | SIMPLE_METHOD : tactic -> method ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1937 | SIMPLE_METHOD' : (int -> tactic) -> method ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1938 | SIMPLE_METHOD'' : ((int -> tactic) -> tactic) -> (int -> tactic) -> method ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1939 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1940 | RAW_METHOD_CASES : (thm list -> cases_tactic) -> method ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1941 | METHOD_CASES : (thm list -> cases_tactic) -> method ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1942 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1943 | 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 | 1944 | the tactic to the current goal state. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1945 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1946 |   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 | 1947 |   \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 | 1948 | goal state. | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1949 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1950 |   \texttt{METHOD} is similar but also first applies
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1951 |   \texttt{Goal.conjunction\_tac} to all subgoals.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1952 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1953 |   \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 | 1954 |   applies \texttt{tacf}.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1955 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1956 |   \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 | 1957 |   applies \texttt{tacf} to subgoal 1.
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1958 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1959 |   \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 | 1960 |   \texttt{quant}, which may be, for example,
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1961 |   \texttt{ALLGOALS} (all subgoals),
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1962 |   \texttt{TRYALL} (try all subgoals, failure is OK),
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1963 |   \texttt{FIRSTGOAL} (try subgoals until it succeeds once), 
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1964 |   \texttt{(fn tacf => tacf 4)} (subgoal 4), etc
 | 
| 16 | 1965 |   (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 | 1966 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1967 | 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 | 1968 |   \begin{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1969 | Method.add_method : | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1970 | (bstring * (src -> Proof.context -> method) * string) -> theory trf ; | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1971 | ( * | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1972 | * ) | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1973 |   \end{verbatim}
 | 
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1974 | 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 | 1975 | 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 | 1976 | (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 | 1977 | |
| 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1978 |   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 | 1979 |   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 | 1980 |   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 | 1981 |   \texttt{Method.add\_method} to add a number of methods.
 | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 1982 | *} | 
| 4 
2a69b119cdee
added verbatim the notes by Jeremy
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1983 | |
| 75 | 1984 | (*>*) | 
| 220 | 1985 | end |