| author | Christian Urban <urbanc@in.tum.de> | 
| Wed, 15 Feb 2012 14:47:46 +0000 | |
| changeset 511 | 386375b7a22a | 
| parent 473 | fea1b7ce5c4a | 
| child 514 | 7e25716c3744 | 
| 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 | |
| 414 | 15 | chapter {* Parsing\label{chp:parsing} *}
 | 
| 4 
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 {*
 | 
| 421 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 18 |   \begin{flushright}
 | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 19 |   {\em An important principle underlying the success and popularity of Unix\\ is
 | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 20 | the philosophy of building on the work of others.} \\[1ex] | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 21 | Linus Torwalds in the email exchange\\ with Andrew S.~Tannenbaum | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 22 |   \end{flushright}
 | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 23 | |
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 24 | |
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 25 |   Isabelle distinguishes between \emph{outer} and \emph{inner}
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 26 |   syntax. Commands, such as \isacommand{definition}, \isacommand{inductive}
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 27 | 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 | 28 | 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 | 29 | general and sophisticated algorithm, which is driven by priority | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 30 | 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 | 31 | combinators. These combinators are a well-established technique for parsing, | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 32 | 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 | 33 |   \cite{paulson-ml2}.  Isabelle developers are usually concerned with writing
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 34 | 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 | 35 | calling methods with specific arguments. | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 36 | |
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 37 |   \begin{readmore}
 | 
| 236 | 38 | 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 | 39 | parts: The first part consists of a collection of generic parser combinators | 
| 236 | 40 |   defined in the structure @{ML_struct Scan} in the file @{ML_file
 | 
| 41 | "Pure/General/scan.ML"}. The second part of the library consists of | |
| 42 | combinators for dealing with specific token types, which are defined in the | |
| 426 | 43 |   structure @{ML_struct Parse} in the file @{ML_file
 | 
| 424 | 44 | "Pure/Isar/parse.ML"}. In addition specific parsers for packages are | 
| 45 |   defined in @{ML_file "Pure/Isar/parse_spec.ML"}. Parsers for method arguments 
 | |
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 46 |   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 | 47 |   \end{readmore}
 | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 48 | |
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 49 | *} | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 50 | |
| 49 | 51 | section {* Building Generic Parsers *}
 | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 52 | |
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 53 | text {*
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 54 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 55 | 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 | 56 |   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 | 57 | ``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 | 58 | 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 | 59 | 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 | 60 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 61 |   @{ML_response [display,gray] 
 | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 62 | "($$ \"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 | 63 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 64 |   @{ML_response [display,gray] 
 | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 65 | "($$ \"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 | 66 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 67 |   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 | 68 |   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 | 69 | example trying to parse | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 70 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 71 |   @{ML_response_fake [display,gray] 
 | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 72 | "($$ \"x\") (Symbol.explode \"world\")" | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 73 | "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 | 74 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 75 |   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 | 76 | the parsing combinators: | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 77 | |
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 78 |   \begin{itemize}
 | 
| 58 | 79 |   \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 | 80 | might be explored. | 
| 58 | 81 |   \item @{text "MORE"} indicates that there is not enough input for the parser. For example 
 | 
| 82 |   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 | 83 |   \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 | 84 |   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 | 85 |   \end{itemize}
 | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 86 | |
| 50 | 87 | However, note that these exceptions are private to the parser and cannot be accessed | 
| 49 | 88 | by the programmer (for example to handle them). | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 89 | |
| 357 
80b56d9b322f
included a comment from Tim Bourke
 Christian Urban <urbanc@in.tum.de> parents: 
346diff
changeset | 90 |   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 | 91 |   structure @{ML_struct Symbol}, instead of the more standard library function
 | 
| 369 | 92 |   @{ML_ind explode in String}, for obtaining an input list for the parser. The reason is
 | 
| 458 
242e81f4d461
updated to post-2011 Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
451diff
changeset | 93 |   that @{ML explode in Symbol} is aware of character
 | 
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 94 |   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 | 95 | Isabelle. To see the difference consider | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 96 | |
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 97 | @{ML_response_fake [display,gray]
 | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 98 | "let | 
| 261 | 99 | val input = \"\<foo> bar\" | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 100 | in | 
| 458 
242e81f4d461
updated to post-2011 Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
451diff
changeset | 101 | (String.explode input, Symbol.explode input) | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 102 | end" | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 103 | "([\"\\\", \"<\", \"f\", \"o\", \"o\", \">\", \" \", \"b\", \"a\", \"r\"], | 
| 261 | 104 | [\"\<foo>\", \" \", \"b\", \"a\", \"r\"])"} | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 105 | |
| 256 
1fb8d62c88a0
added some first index-information
 Christian Urban <urbanc@in.tum.de> parents: 
250diff
changeset | 106 |   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 | 107 |   @{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 | 108 | then parses exactly | 
| 52 | 109 | one item from the input list satisfying this predicate. For example the | 
| 58 | 110 |   following parser either consumes an @{text [quotes] "h"} or a @{text
 | 
| 49 | 111 | [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 | 112 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 113 | @{ML_response [display,gray] 
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 114 | "let | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 115 | 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 | 116 | val input1 = Symbol.explode \"hello\" | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 117 | val input2 = Symbol.explode \"world\" | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 118 | in | 
| 236 | 119 | (hw input1, hw input2) | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 120 | end" | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 121 | "((\"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 | 122 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 123 |   Two parsers can be connected in sequence by using the function @{ML_ind "--" in Scan}. 
 | 
| 220 | 124 |   For example parsing @{text "h"}, @{text "e"} and @{text "l"} (in this 
 | 
| 125 | order) you can achieve by: | |
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 126 | |
| 236 | 127 |   @{ML_response [display,gray] 
 | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 128 | "($$ \"h\" -- $$ \"e\" -- $$ \"l\") (Symbol.explode \"hello\")" | 
| 236 | 129 | "(((\"h\", \"e\"), \"l\"), [\"l\", \"o\"])"} | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 130 | |
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 131 | 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 | 132 | |
| 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 | 133 | 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 | 134 |   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 | 135 | |
| 236 | 136 |   @{ML_response [display,gray] 
 | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 137 | "Scan.this_string \"hell\" (Symbol.explode \"hello\")" | 
| 236 | 138 | "(\"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 | 139 | |
| 256 
1fb8d62c88a0
added some first index-information
 Christian Urban <urbanc@in.tum.de> parents: 
250diff
changeset | 140 | 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 | 141 |   @{ML_ind  "||" in Scan}. The parser @{ML "(p || q)" for p q} returns the
 | 
| 58 | 142 |   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 | 143 |   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 | 144 | |
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 145 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 146 | @{ML_response [display,gray] 
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 147 | "let | 
| 236 | 148 | val hw = $$ \"h\" || $$ \"w\" | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 149 | val input1 = Symbol.explode \"hello\" | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 150 | val input2 = Symbol.explode \"world\" | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 151 | in | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 152 | (hw input1, hw input2) | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 153 | end" | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 154 | "((\"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 | 155 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 156 |   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 | 157 | 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 | 158 | 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 | 159 |   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 | 160 | For example: | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 161 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 162 | @{ML_response [display,gray]
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 163 | "let | 
| 236 | 164 | val just_e = $$ \"h\" |-- $$ \"e\" | 
| 165 | val just_h = $$ \"h\" --| $$ \"e\" | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 166 | val input = Symbol.explode \"hello\" | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 167 | in | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 168 | (just_e input, just_h input) | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 169 | end" | 
| 241 | 170 | "((\"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 | 171 | |
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 172 |   The parser @{ML "Scan.optional p x" for p x} returns the result of the parser 
 | 
| 58 | 173 |   @{text "p"}, if it succeeds; otherwise it returns 
 | 
| 104 | 174 |   the default value @{text "x"}. For example:
 | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 175 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 176 | @{ML_response [display,gray]
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 177 | "let | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 178 | val p = Scan.optional ($$ \"h\") \"x\" | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 179 | val input1 = Symbol.explode \"hello\" | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 180 | val input2 = Symbol.explode \"world\" | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 181 | in | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 182 | (p input1, p input2) | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 183 | end" | 
| 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
38diff
changeset | 184 | "((\"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 | 185 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 186 |   The function @{ML_ind option in Scan} works similarly, except no default value can
 | 
| 50 | 187 |   be given. Instead, the result is wrapped as an @{text "option"}-type. For example:
 | 
| 188 | ||
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 189 | @{ML_response [display,gray]
 | 
| 50 | 190 | "let | 
| 191 | val p = Scan.option ($$ \"h\") | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 192 | val input1 = Symbol.explode \"hello\" | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 193 | val input2 = Symbol.explode \"world\" | 
| 50 | 194 | in | 
| 195 | (p input1, p input2) | |
| 196 | end" "((SOME \"h\", [\"e\", \"l\", \"l\", \"o\"]), (NONE, [\"w\", \"o\", \"r\", \"l\", \"d\"]))"} | |
| 49 | 197 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 198 |   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 | 199 | input unchanged. For example: | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 200 | |
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 201 |   @{ML_response [display,gray]
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 202 | "Scan.ahead (Scan.this_string \"foo\") (Symbol.explode \"foo\")" | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 203 | "(\"foo\", [\"f\", \"o\", \"o\"])"} | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 204 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 205 |   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 | 206 |   during parsing. For example if you want to parse @{text p} immediately 
 | 
| 58 | 207 |   followed by @{text q}, or start a completely different parser @{text r},
 | 
| 104 | 208 | you might write: | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 209 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 210 |   @{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 | 211 | |
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 212 | 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 | 213 |   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 | 214 |   parser above you lose the information that @{text p} should be followed by @{text q}.
 | 
| 220 | 215 |   To see this assume that @{text p} is present in the input, but it is not
 | 
| 216 |   followed by @{text q}. That means @{ML "(p -- q)" for p q} will fail and
 | |
| 217 |   hence the alternative parser @{text r} will be tried. However, in many
 | |
| 236 | 218 |   circumstances this will be the wrong parser for the input ``@{text "p"}-followed-by-something''
 | 
| 220 | 219 | and therefore will also fail. The error message is then caused by the failure | 
| 220 |   of @{text r}, not by the absence of @{text q} in the input. This kind of
 | |
| 221 |   situation can be avoided when using the function @{ML "!!"}.  This function
 | |
| 222 | aborts the whole process of parsing in case of a failure and prints an error | |
| 223 | 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 | 224 | |
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 225 | |
| 472 | 226 |   @{ML [display,gray] "!! (fn _ => fn _ =>\"foo\") ($$ \"h\")"}
 | 
| 227 | *} | |
| 228 | text {*
 | |
| 58 | 229 |   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 | 230 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 231 |   @{ML_response [display,gray] 
 | 
| 472 | 232 | "(!! (fn _ => fn _ => \"foo\") ($$ \"h\")) (Symbol.explode \"hello\")" | 
| 236 | 233 | "(\"h\", [\"e\", \"l\", \"l\", \"o\"])"} | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 234 | |
| 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 | 235 |   but if you invoke it on @{text [quotes] "world"}
 | 
| 472 | 236 | |
| 237 |   @{ML_response_fake [display,gray] "(!! (fn _ => 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 | 238 | "Exception ABORT raised"} | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 239 | |
| 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 | 240 |   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 | 241 | 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 | 242 |   @{ML_ind error in Scan}. For example:
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 243 | |
| 236 | 244 |   @{ML_response_fake [display,gray] 
 | 
| 472 | 245 | "Scan.error (!! (fn _ => fn _ => \"foo\") ($$ \"h\"))" | 
| 236 | 246 | "Exception Error \"foo\" raised"} | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 247 | |
| 426 | 248 |   This ``prefixing'' is usually done by wrappers such as @{ML_ind local_theory in Outer_Syntax} 
 | 
| 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 | 249 |   (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 | 250 | |
| 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 | 251 |   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 | 252 | 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 | 253 |   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 | 254 | *} | 
| 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 255 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 256 | 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 | 257 | let | 
| 236 | 258 | 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 | 259 | in | 
| 472 | 260 | ($$ p -- (!! (fn _ => err_msg) ($$ q))) || ($$ r -- $$ r) | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 261 | end *} | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 262 | |
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 263 | |
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 264 | text {*
 | 
| 220 | 265 | Running this parser with the arguments | 
| 266 |   @{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 | 267 |   the input @{text [quotes] "holle"} 
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 268 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 269 |   @{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 | 270 | "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 | 271 | |
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 272 | produces the correct error message. Running it with | 
| 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.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 | 275 | "((\"w\", \"w\"), [\"o\", \"r\", \"l\", \"d\"])"} | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 276 | |
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 277 | yields the expected parsing. | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 278 | |
| 58 | 279 |   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 | 280 | often as it succeeds. For example: | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 281 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 282 |   @{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 | 283 | "([\"h\", \"h\", \"h\", \"h\"], [\"e\", \"l\", \"l\", \"o\"])"} | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 284 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 285 |   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 | 286 |   @{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 | 287 | succeeds at least once. | 
| 48 
609f9ef73494
fixed FIXME's in fake responses
 Christian Urban <urbanc@in.tum.de> parents: 
47diff
changeset | 288 | |
| 58 | 289 |   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 | 290 |   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 | 291 |   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 | 292 |   @{ML_ind stopper in Symbol}. With them you can write:
 | 
| 49 | 293 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 294 |   @{ML_response [display,gray] "Scan.finite Symbol.stopper (Scan.repeat ($$ \"h\")) (Symbol.explode \"hhhh\")" 
 | 
| 49 | 295 | "([\"h\", \"h\", \"h\", \"h\"], [])"} | 
| 296 | ||
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 297 |   The function @{ML stopper in Symbol} is the ``end-of-input'' indicator for parsing strings;
 | 
| 128 | 298 | 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 | 299 | manually wrapping is often already done by the surrounding infrastructure. | 
| 49 | 300 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 301 |   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 | 302 | 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 | 303 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 304 |   @{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 | 305 | "let | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 306 | 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 | 307 | 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 | 308 | in | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 309 | 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 | 310 | end" | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 311 | "([\"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 | 312 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 313 |   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 | 314 | 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 | 315 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 316 |   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 | 317 | 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 | 318 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 319 |   @{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 | 320 | "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 | 321 | |
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 322 | fails, while | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 323 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 324 |   @{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 | 325 | "(\"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 | 326 | |
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 327 | succeeds. | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 328 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 329 |   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 | 330 | 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 | 331 |   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 | 332 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 333 |   @{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 | 334 | "let | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 335 | 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 | 336 | val input1 = Symbol.explode \"fooooo\" | 
| 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 337 | 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 | 338 | in | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 339 | (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 | 340 | 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 | 341 | end" | 
| 
126646f2aa88
added a para on Scan.unless and an exercise about scanning comments
 Christian Urban <urbanc@in.tum.de> parents: 
54diff
changeset | 342 | "(([\"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 | 343 | ([\"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 | 344 | |
| 256 
1fb8d62c88a0
added some first index-information
 Christian Urban <urbanc@in.tum.de> parents: 
250diff
changeset | 345 | |
| 220 | 346 | 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 | 347 |   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 | 348 |   @{ML "(p >> f)" for p f} runs 
 | 
| 58 | 349 |   first the parser @{text p} and upon successful completion applies the 
 | 
| 350 |   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 | 351 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 352 | @{ML_response [display,gray]
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 353 | "let | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 354 | fun double (x, y) = (x ^ x, y ^ y) | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 355 | val parser = $$ \"h\" -- $$ \"e\" | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 356 | in | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 357 | (parser >> double) (Symbol.explode \"hello\") | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 358 | end" | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 359 | "((\"hh\", \"ee\"), [\"l\", \"l\", \"o\"])"} | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 360 | |
| 104 | 361 | doubles the two parsed input strings; or | 
| 59 | 362 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 363 |   @{ML_response [display,gray] 
 | 
| 59 | 364 | "let | 
| 104 | 365 | 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 | 366 | val input = Symbol.explode \"foo bar foo\" | 
| 59 | 367 | in | 
| 104 | 368 | Scan.finite Symbol.stopper (p >> implode) input | 
| 59 | 369 | end" | 
| 370 | "(\"foo bar foo\",[])"} | |
| 371 | ||
| 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 | 372 | where the single-character strings in the parsed output are transformed | 
| 59 | 373 | 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 | 374 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 375 |   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 | 376 | 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 | 377 | untouched. For example | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 378 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 379 | @{ML_response [display,gray]
 | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
236diff
changeset | 380 | "Scan.lift ($$ \"h\" -- $$ \"e\") (1, Symbol.explode \"hello\")" | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 381 | "((\"h\", \"e\"), (1, [\"l\", \"l\", \"o\"]))"} | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 382 | |
| 390 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 383 |   \footnote{\bf FIXME: In which situations is @{text "lift"} useful? Give examples.} 
 | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 384 | |
| 397 | 385 | Be aware of recursive parsers. Suppose you want to read strings separated by | 
| 386 | commas and by parentheses into a tree datastructure; for example, generating | |
| 387 |   the tree corresponding to the string @{text [quotes] "(A, A), (A, A)"} where
 | |
| 388 |   the @{text "A"}s will be the leaves.  We assume the trees are represented by the
 | |
| 389 | datatype: | |
| 390 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 390 | *} | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 391 | |
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 392 | ML{*datatype tree = 
 | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 393 | Lf of string | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 394 | | Br of tree * tree*} | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 395 | |
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 396 | text {*
 | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 397 | Since nested parentheses should be treated in a meaningful way---for example | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 398 |   the string @{text [quotes] "((A))"} should be read into a single 
 | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 399 | leaf---you might implement the following parser. | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 400 | *} | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 401 | |
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 402 | ML{*fun parse_basic s = 
 | 
| 397 | 403 |   $$ s >> Lf || $$ "(" |-- parse_tree s --| $$ ")"  
 | 
| 404 | ||
| 390 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 405 | and parse_tree s = | 
| 397 | 406 | parse_basic s --| $$ "," -- parse_tree s >> Br || parse_basic s*} | 
| 390 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 407 | |
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 408 | text {*
 | 
| 397 | 409 | This parser corrsponds to the grammar: | 
| 410 | ||
| 411 |   \begin{center}
 | |
| 412 |   \begin{tabular}{lcl}
 | |
| 413 |   @{text "<Basic>"}  & @{text "::="} & @{text "<String> | (<Tree>)"}\\
 | |
| 414 |   @{text "<Tree>"}   & @{text "::="} & @{text "<Basic>, <Tree> | <Basic>"}\\
 | |
| 415 |   \end{tabular}
 | |
| 416 |   \end{center}
 | |
| 417 | ||
| 390 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 418 |   The parameter @{text "s"} is the string over which the tree is parsed. The
 | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 419 |   parser @{ML parse_basic} reads either a leaf or a tree enclosed in
 | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 420 |   parentheses. The parser @{ML parse_tree} reads either a pair of trees
 | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 421 |   separated by a comma, or acts like @{ML parse_basic}. Unfortunately,
 | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 422 | because of the mutual recursion, this parser will immediately run into a | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 423 | loop, even if it is called without any input. For example | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 424 | |
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 425 |   @{ML_response_fake_both [display, gray]
 | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 426 | "parse_tree \"A\"" | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 427 | "*** Exception- TOPLEVEL_ERROR raised"} | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 428 | |
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 429 | raises an exception indicating that the stack limit is reached. Such | 
| 392 
47e5b71c7f6c
modified the typ-pretty-printer and added parser exercise
 Christian Urban <urbanc@in.tum.de> parents: 
391diff
changeset | 430 | looping parser are not useful, because of ML's strict evaluation of | 
| 390 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 431 | arguments. Therefore we need to delay the execution of the | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 432 | parser until an input is given. This can be done by adding the parsed | 
| 397 | 433 | string as an explicit argument. So the parser above should be implemented | 
| 434 | as follows. | |
| 390 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 435 | *} | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 436 | |
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 437 | ML{*fun parse_basic s xs = 
 | 
| 397 | 438 |   ($$ s >> Lf || $$ "(" |-- parse_tree s --| $$ ")") xs 
 | 
| 439 | ||
| 390 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 440 | and parse_tree s xs = | 
| 397 | 441 | (parse_basic s --| $$ "," -- parse_tree s >> Br || parse_basic s) xs*} | 
| 390 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 442 | |
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 443 | text {*
 | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 444 | While the type of the parser is unchanged by the addition, its behaviour | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 445 | changed: with this version of the parser the execution is delayed until | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 446 |   some string is  applied for the argument @{text "xs"}. This gives us 
 | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 447 | exactly the parser what we wanted. An example is as follows: | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 448 | |
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 449 |   @{ML_response [display, gray]
 | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 450 | "let | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 451 | val input = Symbol.explode \"(A,((A))),A\" | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 452 | in | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 453 | Scan.finite Symbol.stopper (parse_tree \"A\") input | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 454 | end" | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 455 | "(Br (Br (Lf \"A\", Lf \"A\"), Lf \"A\"), [])"} | 
| 
8ad407e77ea0
added example by Lukas Bulwahn
 Christian Urban <urbanc@in.tum.de> parents: 
376diff
changeset | 456 | |
| 149 | 457 | |
| 458 |   \begin{exercise}\label{ex:scancmts}
 | |
| 459 | Write a parser that parses an input string so that any comment enclosed | |
| 220 | 460 |   within @{text "(*\<dots>*)"} is replaced by the same comment but enclosed within
 | 
| 149 | 461 |   @{text "(**\<dots>**)"} in the output string. To enclose a string, you can use the
 | 
| 462 |   function @{ML "enclose s1 s2 s" for s1 s2 s} which produces the string @{ML
 | |
| 236 | 463 | "s1 ^ s ^ s2" for s1 s2 s}. Hint: To simplify the task ignore the proper | 
| 464 | nesting of comments. | |
| 149 | 465 |   \end{exercise}
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 466 | *} | 
| 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 467 | |
| 41 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 468 | section {* Parsing Theory Syntax *}
 | 
| 38 
e21b2f888fa2
added a preliminary section about parsing
 Christian Urban <urbanc@in.tum.de> parents: 
16diff
changeset | 469 | |
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 470 | 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 | 471 | Most of the time, however, Isabelle developers have to deal with parsing | 
| 156 | 472 | tokens, not strings. These token parsers have the type: | 
| 128 | 473 | *} | 
| 474 | ||
| 426 | 475 | ML{*type 'a parser = Token.T list -> 'a * Token.T list*}
 | 
| 128 | 476 | |
| 477 | text {*
 | |
| 149 | 478 | The reason for using token parsers is that theory syntax, as well as the | 
| 128 | 479 |   parsers for the arguments of proof methods, use the type @{ML_type
 | 
| 426 | 480 | Token.T}. | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 481 | |
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 482 |   \begin{readmore}
 | 
| 40 
35e1dff0d9bb
more on the parsing section
 Christian Urban <urbanc@in.tum.de> parents: 
39diff
changeset | 483 | The parser functions for the theory syntax are contained in the structure | 
| 426 | 484 |   @{ML_struct Parse} defined in the file @{ML_file  "Pure/Isar/parse.ML"}.
 | 
| 485 |   The definition for tokens is in the file @{ML_file "Pure/Isar/token.ML"}.
 | |
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 486 |   \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 | 487 | |
| 426 | 488 |   The structure @{ML_struct  Token} defines several kinds of tokens (for
 | 
| 489 |   example @{ML_ind Ident in Token} for identifiers, @{ML Keyword in
 | |
| 490 |   Token} for keywords and @{ML_ind Command in Token} for commands). Some
 | |
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 491 | 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 | 492 | how to generate a token list out of a string using the function | 
| 426 | 493 |   @{ML_ind scan in Outer_Syntax}. It is given the argument 
 | 
| 256 
1fb8d62c88a0
added some first index-information
 Christian Urban <urbanc@in.tum.de> parents: 
250diff
changeset | 494 |   @{ML "Position.none"} since,
 | 
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 495 | at the moment, we are not interested in generating precise error | 
| 376 
76b1b679e845
removed comment about compiler bug
 Christian Urban <urbanc@in.tum.de> parents: 
374diff
changeset | 496 | messages. The following code | 
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 497 | |
| 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 | 498 | |
| 426 | 499 | @{ML_response_fake [display,gray] "Outer_Syntax.scan Position.none \"hello world\"" 
 | 
| 50 | 500 | "[Token (\<dots>,(Ident, \"hello\"),\<dots>), | 
| 501 | Token (\<dots>,(Space, \" \"),\<dots>), | |
| 502 | Token (\<dots>,(Ident, \"world\"),\<dots>)]"} | |
| 503 | ||
| 504 | produces three tokens where the first and the last are identifiers, since | |
| 58 | 505 |   @{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 | 506 | 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 | 507 | |
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 508 | We can easily change what is recognised as a keyword with the function | 
| 426 | 509 |   @{ML_ind keyword in Keyword}. For example calling it with 
 | 
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 510 | *} | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 511 | |
| 426 | 512 | ML{*val _ = Keyword.keyword "hello"*}
 | 
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 513 | |
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 514 | text {*
 | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 515 |   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 | 516 | |
| 426 | 517 |   @{ML_response_fake [display,gray] "Outer_Syntax.scan Position.none \"hello world\"" 
 | 
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 518 | "[Token (\<dots>,(Keyword, \"hello\"),\<dots>), | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 519 | Token (\<dots>,(Space, \" \"),\<dots>), | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 520 | Token (\<dots>,(Ident, \"world\"),\<dots>)]"} | 
| 50 | 521 | |
| 241 | 522 | 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 | 523 | to have already been filtered out. So from now on we are going to use the | 
| 426 | 524 |   functions @{ML filter} and @{ML_ind is_proper in Token} to do this. 
 | 
| 256 
1fb8d62c88a0
added some first index-information
 Christian Urban <urbanc@in.tum.de> parents: 
250diff
changeset | 525 | 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 | 526 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 527 | @{ML_response_fake [display,gray]
 | 
| 50 | 528 | "let | 
| 426 | 529 | val input = Outer_Syntax.scan Position.none \"hello world\" | 
| 50 | 530 | in | 
| 426 | 531 | filter Token.is_proper input | 
| 50 | 532 | end" | 
| 533 | "[Token (\<dots>,(Ident, \"hello\"), \<dots>), Token (\<dots>,(Ident, \"world\"), \<dots>)]"} | |
| 534 | ||
| 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 | 535 | For convenience we define the function: | 
| 50 | 536 | *} | 
| 537 | ||
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 538 | ML{*fun filtered_input str = 
 | 
| 426 | 539 | filter Token.is_proper (Outer_Syntax.scan Position.none str) *} | 
| 50 | 540 | |
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 541 | 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 | 542 | 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 | 543 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 544 | @{ML_response_fake [display,gray] 
 | 
| 50 | 545 | "filtered_input \"inductive | for\"" | 
| 546 | "[Token (\<dots>,(Command, \"inductive\"),\<dots>), | |
| 547 | Token (\<dots>,(Keyword, \"|\"),\<dots>), | |
| 548 | 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 | 549 | |
| 221 | 550 | you obtain a list consisting of only one command and two keyword tokens. | 
| 241 | 551 | If you want to see which keywords and commands are currently known to Isabelle, | 
| 449 | 552 |   use the function @{ML_ind get_lexicons in Keyword}:
 | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 553 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 554 | @{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 | 555 | "let | 
| 426 | 556 | val (keywords, commands) = Keyword.get_lexicons () | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 557 | in | 
| 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
44diff
changeset | 558 | (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 | 559 | end" | 
| 132 | 560 | "([\"}\", \"{\", \<dots>], [\"\<rightleftharpoons>\", \"\<leftharpoondown>\", \<dots>])"}
 | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 561 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 562 |   You might have to adjust the @{ML_ind print_depth} in order to
 | 
| 241 | 563 | see the complete list. | 
| 564 | ||
| 426 | 565 |   The parser @{ML_ind "$$$" in Parse} parses a single keyword. For example:
 | 
| 50 | 566 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 567 | @{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 | 568 | "let | 
| 50 | 569 | val input1 = filtered_input \"where for\" | 
| 570 | 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 | 571 | in | 
| 426 | 572 | (Parse.$$$ \"where\" input1, Parse.$$$ \"|\" input2) | 
| 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 | 573 | end" | 
| 128 | 574 | "((\"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 | 575 | |
| 426 | 576 |   Any non-keyword string can be parsed with the function @{ML_ind reserved in Parse}.
 | 
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 577 | For example: | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 578 | |
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 579 |   @{ML_response [display,gray]
 | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 580 | "let | 
| 426 | 581 | val p = Parse.reserved \"bar\" | 
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 582 | val input = filtered_input \"bar\" | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 583 | in | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 584 | p input | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 585 | end" | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 586 | "(\"bar\",[])"} | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 587 | |
| 344 
83d5bca38bec
added structures in the index
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 588 |   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 | 589 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 590 | @{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 | 591 | "let | 
| 50 | 592 | 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 | 593 | in | 
| 426 | 594 | (Parse.$$$ \"|\" -- Parse.$$$ \"in\") input | 
| 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 | 595 | end" | 
| 183 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
181diff
changeset | 596 | "((\"|\", \"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 | 597 | |
| 426 | 598 |   The parser @{ML "Parse.enum s p" for s p} parses a possibly empty 
 | 
| 58 | 599 |   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 | 600 |   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 | 601 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 602 | @{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 | 603 | "let | 
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 604 | 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 | 605 | in | 
| 426 | 606 | (Parse.enum \"|\" (Parse.$$$ \"in\")) input | 
| 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 | 607 | end" | 
| 183 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
181diff
changeset | 608 | "([\"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 | 609 | |
| 426 | 610 |   The function @{ML_ind enum1 in Parse} works similarly, except that the
 | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 611 |   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 | 612 | [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 | 613 |   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 | 614 | "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 | 615 |   wrapper @{ML Scan.finite}. This time, however, we have to use the
 | 
| 426 | 616 |   ``stopper-token'' @{ML Token.stopper}. We can write:
 | 
| 49 | 617 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 618 | @{ML_response [display,gray]
 | 
| 49 | 619 | "let | 
| 50 | 620 | val input = filtered_input \"in | in | in\" | 
| 426 | 621 | val p = Parse.enum \"|\" (Parse.$$$ \"in\") | 
| 49 | 622 | in | 
| 426 | 623 | Scan.finite Token.stopper p input | 
| 49 | 624 | end" | 
| 183 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
181diff
changeset | 625 | "([\"in\", \"in\", \"in\"], [])"} | 
| 49 | 626 | |
| 75 | 627 | 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 | 628 | *} | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 629 | |
| 426 | 630 | ML{*fun parse p input = Scan.finite Token.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 | 631 | |
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 632 | text {*
 | 
| 426 | 633 |   The function @{ML_ind "!!!" in Parse} can be used to force termination
 | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 634 |   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 | 635 |   section). A difference, however, is that the error message of @{ML
 | 
| 426 | 636 |   "Parse.!!!"} is fixed to be @{text [quotes] "Outer syntax error"}
 | 
| 221 | 637 | together with a relatively precise description of the failure. For example: | 
| 49 | 638 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 639 | @{ML_response_fake [display,gray]
 | 
| 49 | 640 | "let | 
| 50 | 641 | val input = filtered_input \"in |\" | 
| 426 | 642 | val parse_bar_then_in = Parse.$$$ \"|\" -- Parse.$$$ \"in\" | 
| 49 | 643 | in | 
| 426 | 644 | parse (Parse.!!! parse_bar_then_in) input | 
| 49 | 645 | end" | 
| 646 | "Exception ERROR \"Outer syntax error: keyword \"|\" expected, | |
| 647 | but keyword in was found\" raised" | |
| 648 | } | |
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 649 | |
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 650 |   \begin{exercise} (FIXME)
 | 
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 651 |   A type-identifier, for example @{typ "'a"}, is a token of 
 | 
| 426 | 652 |   kind @{ML_ind Keyword in Token}. It can be parsed using 
 | 
| 653 |   the function @{ML type_ident in Parse}.
 | |
| 53 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 654 |   \end{exercise}
 | 
| 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 655 | |
| 104 | 656 | (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 | 657 | |
| 125 | 658 | Whenever there is a possibility that the processing of user input can fail, | 
| 221 | 659 | it is a good idea to give all available information about where the error | 
| 220 | 660 | 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 | 661 | 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 | 662 |   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 | 663 | *} | 
| 
b11653b11bd3
further progress on the parsing section and tuning on the antiqu's
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 664 | |
| 125 | 665 | ML{*fun filtered_input' str = 
 | 
| 426 | 666 | filter Token.is_proper (Outer_Syntax.scan (Position.line 7) str) *} | 
| 49 | 667 | |
| 668 | text {*
 | |
| 125 | 669 | where we pretend the parsed string starts on line 7. An example is | 
| 49 | 670 | |
| 125 | 671 | @{ML_response_fake [display,gray]
 | 
| 672 | "filtered_input' \"foo \\n bar\"" | |
| 673 | "[Token ((\"foo\", ({line=7, end_line=7}, {line=7})), (Ident, \"foo\"), \<dots>),
 | |
| 674 |  Token ((\"bar\", ({line=8, end_line=8}, {line=8})), (Ident, \"bar\"), \<dots>)]"}
 | |
| 675 | ||
| 676 |   in which the @{text [quotes] "\\n"} causes the second token to be in 
 | |
| 677 | line 8. | |
| 678 | ||
| 426 | 679 |   By using the parser @{ML position in Parse} you can access the token 
 | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 680 | position and return it as part of the parser result. For example | 
| 125 | 681 | |
| 682 | @{ML_response_fake [display,gray]
 | |
| 683 | "let | |
| 241 | 684 | val input = filtered_input' \"where\" | 
| 125 | 685 | in | 
| 426 | 686 | parse (Parse.position (Parse.$$$ \"where\")) input | 
| 125 | 687 | end" | 
| 688 | "((\"where\", {line=7, end_line=7}), [])"}
 | |
| 689 | ||
| 690 |   \begin{readmore}
 | |
| 691 | The functions related to positions are implemented in the file | |
| 692 |   @{ML_file "Pure/General/position.ML"}.
 | |
| 693 |   \end{readmore}
 | |
| 49 | 694 | |
| 391 | 695 |   \begin{exercise}\label{ex:contextfree}
 | 
| 696 | Write a parser for the context-free grammar representing arithmetic | |
| 697 | expressions with addition and multiplication. As usual, multiplication | |
| 698 | binds stronger than addition, and both of them nest to the right. | |
| 699 | The context-free grammar is defined as: | |
| 700 | ||
| 701 |   \begin{center}
 | |
| 702 |   \begin{tabular}{lcl}
 | |
| 703 |   @{text "<Basic>"}  & @{text "::="} & @{text "<Number> | (<Expr>)"}\\
 | |
| 704 |   @{text "<Factor>"} & @{text "::="} & @{text "<Basic> * <Factor> | <Basic>"}\\
 | |
| 705 |   @{text "<Expr>"}   & @{text "::="} & @{text "<Factor> + <Expr>  | <Factor>"}\\
 | |
| 706 |   \end{tabular}
 | |
| 707 |   \end{center}
 | |
| 708 | ||
| 709 | Hint: Be careful with recursive parsers. | |
| 710 |   \end{exercise}
 | |
| 49 | 711 | *} | 
| 712 | ||
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 713 | section {* Parsers for ML-Code (TBD) *}
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 714 | |
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 715 | text {*
 | 
| 426 | 716 |   @{ML_ind ML_source in Parse}
 | 
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 717 | *} | 
| 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 718 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 719 | section {* Context Parser (TBD) *}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 720 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 721 | text {*
 | 
| 326 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 722 |   @{ML_ind Args.context}
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 723 | *} | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 724 | (* | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 725 | ML {*
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 726 | let | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 727 | val parser = Args.context -- Scan.lift Args.name_source | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 728 | |
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 729 | fun term_pat (ctxt, str) = | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 730 | str |> Syntax.read_prop ctxt | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 731 | in | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 732 |   (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 | 733 | |> fst | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 734 | end | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 735 | *} | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 736 | *) | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 737 | |
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 738 | text {*
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 739 |   @{ML_ind Args.context}
 | 
| 
f76135c6c527
more work on the tutorial
 Christian Urban <urbanc@in.tum.de> parents: 
324diff
changeset | 740 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 741 |   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 | 742 | *} | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 743 | |
| 207 | 744 | section {* Argument and Attribute Parsers (TBD) *}
 | 
| 745 | ||
| 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 | 746 | section {* Parsing Inner Syntax *}
 | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 747 | |
| 125 | 748 | text {*
 | 
| 749 | There is usually no need to write your own parser for parsing inner syntax, that is | |
| 285 | 750 | for terms and types: you can just call the predefined parsers. Terms can | 
| 426 | 751 |   be parsed using the function @{ML_ind term in Parse}. For example:
 | 
| 125 | 752 | |
| 753 | @{ML_response [display,gray]
 | |
| 754 | "let | |
| 426 | 755 | val input = Outer_Syntax.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 | 756 | in | 
| 426 | 757 | Parse.term input | 
| 125 | 758 | end" | 
| 759 | "(\"\\^E\\^Ftoken\\^Efoo\\^E\\^F\\^E\", [])"} | |
| 760 | ||
| 426 | 761 |   The function @{ML_ind prop in Parse} 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 | 762 | 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 | 763 | 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 | 764 |   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 | 765 | |
| 445 | 766 |   @{ML_response_fake [display,gray]
 | 
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 767 | "YXML.parse \"\\^E\\^Ftoken\\^Efoo\\^E\\^F\\^E\"" | 
| 445 | 768 | "Text \"\\^E\\^Ftoken\\^Efoo\\^E\\^F\\^E\""} | 
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 769 | |
| 149 | 770 | The result of the decoding is an XML-tree. You can see better what is going on if | 
| 131 | 771 |   you replace @{ML Position.none} by @{ML "Position.line 42"}, say:
 | 
| 101 | 772 | |
| 445 | 773 | @{ML_response_fake [display,gray]
 | 
| 125 | 774 | "let | 
| 426 | 775 | val input = Outer_Syntax.scan (Position.line 42) \"foo\" | 
| 125 | 776 | in | 
| 426 | 777 | YXML.parse (fst (Parse.term input)) | 
| 125 | 778 | end" | 
| 445 | 779 | "Elem (\"token\", [(\"line\", \"42\"), (\"end_line\", \"42\")], [XML.Text \"foo\"])"} | 
| 780 | ||
| 149 | 781 | The positional information is stored as part of an XML-tree so that code | 
| 782 | called later on will be able to give more precise error messages. | |
| 125 | 783 | |
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 784 |   \begin{readmore}
 | 
| 128 | 785 | The functions to do with input and output of XML and YXML are defined | 
| 473 
fea1b7ce5c4a
this version works with Isabelle2011-1
 Christian Urban <urbanc@in.tum.de> parents: 
472diff
changeset | 786 |   in @{ML_file "Pure/PIDE/xml.ML"} and @{ML_file "Pure/PIDE/yxml.ML"}.
 | 
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 787 |   \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 | 788 | |
| 361 | 789 | FIXME: | 
| 790 |   @{ML_ind parse_term in Syntax} @{ML_ind check_term in Syntax}
 | |
| 791 |   @{ML_ind parse_typ in Syntax} @{ML_ind check_typ in Syntax}
 | |
| 374 | 792 |   @{ML_ind read_term in Syntax} @{ML_ind read_term in Syntax}
 | 
| 793 | ||
| 361 | 794 | |
| 125 | 795 | *} | 
| 101 | 796 | |
| 116 
c9ff326e3ce5
more changes to the package chapter
 Christian Urban <urbanc@in.tum.de> parents: 
114diff
changeset | 797 | section {* Parsing Specifications\label{sec:parsingspecs} *}
 | 
| 101 | 798 | |
| 799 | text {*
 | |
| 121 | 800 | There are a number of special purpose parsers that help with parsing | 
| 156 | 801 | specifications of function definitions, inductive predicates and so on. In | 
| 220 | 802 |   Chapter~\ref{chp:package}, for example, we will need to parse specifications
 | 
| 121 | 803 | for inductive predicates of the form: | 
| 804 | *} | |
| 101 | 805 | |
| 451 
fc074e669f9f
disabled foobar_prove; updated to new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
449diff
changeset | 806 | |
| 121 | 807 | simple_inductive | 
| 808 | even and odd | |
| 809 | where | |
| 810 | even0: "even 0" | |
| 811 | | evenS: "odd n \<Longrightarrow> even (Suc n)" | |
| 812 | | oddS: "even n \<Longrightarrow> odd (Suc n)" | |
| 101 | 813 | |
| 814 | text {*
 | |
| 121 | 815 | For this we are going to use the parser: | 
| 101 | 816 | *} | 
| 817 | ||
| 121 | 818 | ML %linenosgray{*val spec_parser = 
 | 
| 426 | 819 | Parse.fixes -- | 
| 126 | 820 | Scan.optional | 
| 426 | 821 | (Parse.$$$ "where" |-- | 
| 822 | Parse.!!! | |
| 823 | (Parse.enum1 "|" | |
| 824 | (Parse_Spec.opt_thm_name ":" -- Parse.prop))) []*} | |
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: 
116diff
changeset | 825 | |
| 101 | 826 | text {*
 | 
| 241 | 827 | Note that the parser must not parse the keyword \simpleinductive, even if it is | 
| 126 | 828 | meant to process definitions as shown above. The parser of the keyword | 
| 128 | 829 |   will be given by the infrastructure that will eventually call @{ML spec_parser}.
 | 
| 126 | 830 | |
| 831 | ||
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 832 | To see what the parser returns, let us parse the string corresponding to the | 
| 121 | 833 |   definition of @{term even} and @{term odd}:
 | 
| 834 | ||
| 101 | 835 | @{ML_response [display,gray]
 | 
| 836 | "let | |
| 837 | val input = filtered_input | |
| 838 | (\"even and odd \" ^ | |
| 839 | \"where \" ^ | |
| 840 | \" even0[intro]: \\\"even 0\\\" \" ^ | |
| 841 | \"| evenS[intro]: \\\"odd n \<Longrightarrow> even (Suc n)\\\" \" ^ | |
| 842 | \"| oddS[intro]: \\\"even n \<Longrightarrow> odd (Suc n)\\\"\") | |
| 843 | in | |
| 120 
c39f83d8daeb
some polishing; split up the file External Solver into two
 Christian Urban <urbanc@in.tum.de> parents: 
116diff
changeset | 844 | parse spec_parser input | 
| 101 | 845 | end" | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
183diff
changeset | 846 | "(([(even, NONE, NoSyn), (odd, NONE, NoSyn)], | 
| 101 | 847 | [((even0,\<dots>), \"\\^E\\^Ftoken\\^Eeven 0\\^E\\^F\\^E\"), | 
| 848 | ((evenS,\<dots>), \"\\^E\\^Ftoken\\^Eodd n \<Longrightarrow> even (Suc n)\\^E\\^F\\^E\"), | |
| 849 | ((oddS,\<dots>), \"\\^E\\^Ftoken\\^Eeven n \<Longrightarrow> odd (Suc n)\\^E\\^F\\^E\")]), [])"} | |
| 121 | 850 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
183diff
changeset | 851 | 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 | 852 | 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 | 853 | rules where every rule has optionally a name and an attribute. | 
| 121 | 854 | |
| 426 | 855 |   The function @{ML_ind "fixes" in Parse} 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 | 856 |   \isacommand{and}-separated 
 | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 857 | list of variables that can include optional type annotations and syntax translations. | 
| 121 | 858 |   For example:\footnote{Note that in the code we need to write 
 | 
| 859 |   @{text "\\\"int \<Rightarrow> bool\\\""} in order to properly escape the double quotes
 | |
| 860 | in the compound type.} | |
| 861 | ||
| 862 | @{ML_response [display,gray]
 | |
| 863 | "let | |
| 864 | val input = filtered_input | |
| 865 | \"foo::\\\"int \<Rightarrow> bool\\\" and bar::nat (\\\"BAR\\\" 100) and blonk\" | |
| 866 | in | |
| 426 | 867 | parse Parse.fixes input | 
| 121 | 868 | end" | 
| 869 | "([(foo, SOME \"\\^E\\^Ftoken\\^Eint \<Rightarrow> bool\\^E\\^F\\^E\", NoSyn), | |
| 870 | (bar, SOME \"\\^E\\^Ftoken\\^Enat\\^E\\^F\\^E\", Mixfix (\"BAR\", [], 100)), | |
| 871 | (blonk, NONE, NoSyn)],[])"} | |
| 50 | 872 | *} | 
| 873 | ||
| 121 | 874 | text {*
 | 
| 156 | 875 |   Whenever types are given, they are stored in the @{ML SOME}s. The types are
 | 
| 876 | not yet used to type the variables: this must be done by type-inference later | |
| 149 | 877 | on. Since types are part of the inner syntax they are strings with some | 
| 241 | 878 | encoded information (see previous section). If a mixfix-syntax is | 
| 369 | 879 | present for a variable, then it is stored in the | 
| 371 
e6f583366779
solved problem with mixfix.
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 880 |   @{ML Mixfix} data structure; no syntax translation is indicated by @{ML_ind NoSyn in Syntax}.
 | 
| 121 | 881 | |
| 882 |   \begin{readmore}
 | |
| 371 
e6f583366779
solved problem with mixfix.
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 883 | The data structure for mixfix annotations are implemented in | 
| 
e6f583366779
solved problem with mixfix.
 Christian Urban <urbanc@in.tum.de> parents: 
369diff
changeset | 884 |   @{ML_file "Pure/Syntax/mixfix.ML"} and @{ML_file "Pure/Syntax/syntax.ML"}.
 | 
| 121 | 885 |   \end{readmore}
 | 
| 886 | ||
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
183diff
changeset | 887 |   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 | 888 | 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 | 889 | such as rule names and attributes. The introduction rules are propositions | 
| 426 | 890 |   parsed by @{ML_ind prop in Parse}. However, they can include an optional
 | 
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 891 | theorem name plus some attributes. For example | 
| 121 | 892 | |
| 893 | @{ML_response [display,gray] "let 
 | |
| 894 | val input = filtered_input \"foo_lemma[intro,dest!]:\" | |
| 426 | 895 | val ((name, attrib), _) = parse (Parse_Spec.thm_name \":\") input | 
| 121 | 896 | in | 
| 897 | (name, map Args.dest_src attrib) | |
| 898 | end" "(foo_lemma, [((\"intro\", []), \<dots>), ((\"dest\", [\<dots>]), \<dots>)])"} | |
| 899 | ||
| 426 | 900 |   The function @{ML_ind opt_thm_name in Parse_Spec} is the ``optional'' variant of
 | 
| 901 |   @{ML_ind thm_name in Parse_Spec}. Theorem names can contain attributes. The name 
 | |
| 131 | 902 |   has to end with @{text [quotes] ":"}---see the argument of 
 | 
| 426 | 903 |   the function @{ML Parse_Spec.opt_thm_name} in Line 7.
 | 
| 121 | 904 | |
| 905 |   \begin{readmore}
 | |
| 906 |   Attributes and arguments are implemented in the files @{ML_file "Pure/Isar/attrib.ML"} 
 | |
| 907 |   and @{ML_file "Pure/Isar/args.ML"}.
 | |
| 908 |   \end{readmore}
 | |
| 101 | 909 | *} | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 910 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 911 | text_raw {*
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 912 |   \begin{exercise}
 | 
| 426 | 913 |   Have a look at how the parser @{ML Parse_Spec.where_alt_specs} is implemented
 | 
| 424 | 914 |   in file @{ML_file "Pure/Isar/parse_spec.ML"}. This parser corresponds
 | 
| 207 | 915 | to the ``where-part'' of the introduction rules given above. Below | 
| 426 | 916 |   we paraphrase the code of @{ML_ind where_alt_specs in Parse_Spec} adapted to our
 | 
| 207 | 917 | purposes. | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 918 |   \begin{isabelle}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 919 | *} | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 920 | ML %linenosgray{*val spec_parser' = 
 | 
| 426 | 921 | Parse.fixes -- | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 922 | Scan.optional | 
| 426 | 923 | (Parse.$$$ "where" |-- | 
| 924 | Parse.!!! | |
| 925 | (Parse.enum1 "|" | |
| 926 | ((Parse_Spec.opt_thm_name ":" -- Parse.prop) --| | |
| 927 | Scan.option (Scan.ahead (Parse.name || | |
| 928 | Parse.$$$ "[") -- | |
| 929 | Parse.!!! (Parse.$$$ "|"))))) [] *} | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 930 | text_raw {*
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 931 |   \end{isabelle}
 | 
| 284 | 932 | Both parsers accept the same input% that's not true: | 
| 933 | % spec_parser accepts input that is refuted by spec_parser' | |
| 934 | , but if you look closely, you can notice | |
| 207 | 935 |   an additional  ``tail'' (Lines 8 to 10) in @{ML spec_parser'}. What is the purpose of 
 | 
| 936 | this additional ``tail''? | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 937 |   \end{exercise}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 938 | *} | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 939 | |
| 229 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 940 | text {*
 | 
| 426 | 941 |   (FIXME: @{ML Parse.type_args}, @{ML Parse.typ}, @{ML Parse.opt_mixfix})
 | 
| 229 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 942 | *} | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 943 | |
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 944 | |
| 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 | 945 | 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 | 946 | |
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 947 | text {*
 | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 948 | 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 | 949 | need to be implemented. While this is not difficult on the ML-level, | 
| 66 | 950 | 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 | 951 | 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 | 952 | will explain in this section. | 
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 953 | |
| 74 | 954 | To keep things simple, let us start with a ``silly'' command that does nothing | 
| 955 |   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 | 956 | defined as: | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 957 | *} | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 958 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 959 | ML{*let
 | 
| 449 | 960 | val do_nothing = Scan.succeed (Local_Theory.background_theory I) | 
| 426 | 961 | val kind = Keyword.thy_decl | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 962 | in | 
| 426 | 963 | Outer_Syntax.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 | 964 | end *} | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 965 | |
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 966 | text {*
 | 
| 426 | 967 |   The crucial function @{ML_ind local_theory in Outer_Syntax} 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 | 968 | 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 | 969 | parser producing a local theory transition (its purpose will also explained | 
| 66 | 970 | later). | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 971 | |
| 101 | 972 | 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 | 973 | 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 | 974 |   recognise \isacommand{foobar} as a command. Such a keyword file can be
 | 
| 74 | 975 | 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 | 976 | |
| 74 | 977 |   @{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 | 978 | |
| 74 | 979 |   The option @{text "-k foobar"} indicates which postfix the name of the keyword file 
 | 
| 80 | 980 |   will be assigned. In the case above the file will be named @{text
 | 
| 86 | 981 | "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 | 982 | present (in order to extract the keywords from them). To generate these log | 
| 101 | 983 | 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 | 984 |   @{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 | 985 | complete code. | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 986 | |
| 66 | 987 | |
| 988 | %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% | |
| 989 |   \begin{figure}[t]
 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 990 |   \begin{graybox}\small
 | 
| 66 | 991 |   \isacommand{theory}~@{text Command}\\
 | 
| 992 |   \isacommand{imports}~@{text Main}\\
 | |
| 993 |   \isacommand{begin}\\
 | |
| 85 
b02904872d6b
better handling of {* and *}
 Christian Urban <urbanc@in.tum.de> parents: 
81diff
changeset | 994 |   \isacommand{ML}~@{text "\<verbopen>"}\\
 | 
| 66 | 995 |   @{ML
 | 
| 996 | "let | |
| 449 | 997 | val do_nothing = Scan.succeed (Local_Theory.background_theory I) | 
| 426 | 998 | val kind = Keyword.thy_decl | 
| 66 | 999 | in | 
| 426 | 1000 | Outer_Syntax.local_theory \"foobar\" \"description of foobar\" kind do_nothing | 
| 66 | 1001 | end"}\\ | 
| 85 
b02904872d6b
better handling of {* and *}
 Christian Urban <urbanc@in.tum.de> parents: 
81diff
changeset | 1002 |   @{text "\<verbclose>"}\\
 | 
| 66 | 1003 |   \isacommand{end}
 | 
| 80 | 1004 |   \end{graybox}
 | 
| 241 | 1005 |   \caption{This file can be used to generate a log file. This log file in turn can
 | 
| 1006 |   be used to generate a keyword file containing the command \isacommand{foobar}.
 | |
| 1007 |   \label{fig:commandtheory}}
 | |
| 66 | 1008 |   \end{figure}
 | 
| 1009 | %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% | |
| 1010 | ||
| 75 | 1011 | 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 | 1012 |   @{text "Pure"}, @{text "HOL"} and @{text "Pure-ProofGeneral"}, as well as
 | 
| 75 | 1013 |   the log file for the theory @{text "Command.thy"}, which contains the new
 | 
| 1014 |   \isacommand{foobar}-command. If you target other logics besides HOL, such
 | |
| 74 | 1015 | as Nominal or ZF, then you need to adapt the log files appropriately. | 
| 104 | 1016 | |
| 74 | 1017 |   @{text Pure} and @{text HOL} are usually compiled during the installation of
 | 
| 1018 | Isabelle. So log files for them should be already available. If not, then | |
| 75 | 1019 | 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 | 1020 | distribution. | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1021 | |
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1022 |   @{text [display] 
 | 
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1023 | "$ ./build -m \"Pure\" | 
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1024 | $ ./build -m \"HOL\""} | 
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1025 | |
| 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 | 1026 |   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 | 1027 | |
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1028 |   @{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 | 1029 | |
| 101 | 1030 |   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 | 1031 | with: | 
| 66 | 1032 | |
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1033 |   @{text [display] "$ isabelle mkdir FoobarCommand"}
 | 
| 66 | 1034 | |
| 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 | 1035 | This generates a directory containing the files: | 
| 66 | 1036 | |
| 1037 |   @{text [display] 
 | |
| 1038 | "./IsaMakefile | |
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1039 | ./FoobarCommand/ROOT.ML | 
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1040 | ./FoobarCommand/document | 
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1041 | ./FoobarCommand/document/root.tex"} | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1042 | |
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1043 | |
| 101 | 1044 |   You need to copy the file @{text "Command.thy"} into the directory @{text "FoobarCommand"}
 | 
| 66 | 1045 | and add the line | 
| 1046 | ||
| 207 | 1047 |   @{text [display] "no_document use_thy \"Command\";"} 
 | 
| 66 | 1048 | |
| 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 | 1049 |   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 | 1050 | |
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1051 |   @{text [display] "$ isabelle make"}
 | 
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1052 | |
| 101 | 1053 | If the compilation succeeds, you have finally created all the necessary log files. | 
| 1054 | 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 | 1055 | |
| 241 | 1056 |   @{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 | 1057 | |
| 74 | 1058 | or something similar depending on your Isabelle distribution and architecture. | 
| 1059 | One quick way to assign a shell variable to this directory is by typing | |
| 66 | 1060 | |
| 1061 |   @{text [display] "$ ISABELLE_LOGS=\"$(isabelle getenv -b ISABELLE_OUTPUT)\"/log"}
 | |
| 1062 | ||
| 156 | 1063 |   on the Unix prompt. If you now type @{text "ls $ISABELLE_LOGS"}, then the 
 | 
| 128 | 1064 | 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 | 1065 | |
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1066 |   @{text [display] 
 | 
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1067 | "Pure.gz | 
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1068 | HOL.gz | 
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1069 | Pure-ProofGeneral.gz | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1070 | HOL-FoobarCommand.gz"} | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1071 | |
| 101 | 1072 | From them you can create the keyword files. Assuming the name | 
| 75 | 1073 |   of the directory is in  @{text "$ISABELLE_LOGS"},
 | 
| 74 | 1074 | 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 | 1075 | |
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1076 | @{text [display]
 | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1077 | "$ isabelle keywords -k foobar | 
| 80 | 1078 |    $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 | 1079 | |
| 80 | 1080 |   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 | 1081 |   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 | 1082 |   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 | 1083 |   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 | 1084 | "~/.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 | 1085 |   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 | 1086 | |
| 80 | 1087 | |
| 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 | 1088 |   @{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 | 1089 | |
| 101 | 1090 |   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 | 1091 |   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 | 1092 | *} | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1093 | |
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1094 | foobar | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1095 | |
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1096 | text {* 
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1097 | 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 | 1098 | 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 | 1099 | 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 | 1100 | other new command, and also any new keyword that is introduced with | 
| 426 | 1101 |   the function @{ML_ind keyword in Keyword}. For example:
 | 
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 1102 | *} | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1103 | |
| 426 | 1104 | ML{*val _ = Keyword.keyword "blink" *}
 | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1105 | |
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 1106 | text {*
 | 
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1107 |   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 | 1108 | 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 | 1109 | 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 | 1110 | |
| 75 | 1111 | The crucial part of a command is the function that determines the behaviour | 
| 1112 | 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 | 1113 |   because of @{ML_ind succeed in Scan} does not parse any argument, but immediately
 | 
| 449 | 1114 |   returns the simple function @{ML "Local_Theory.background_theory I"}. We can
 | 
| 75 | 1115 | replace this code by a function that first parses a proposition (using the | 
| 426 | 1116 |   parser @{ML Parse.prop}), then prints out the tracing
 | 
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 1117 |   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 | 1118 | 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 | 1119 | *} | 
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1120 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 1121 | ML{*let
 | 
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 1122 | fun trace_prop str = | 
| 449 | 1123 | Local_Theory.background_theory (fn ctxt => (tracing str; ctxt)) | 
| 75 | 1124 | |
| 426 | 1125 | val kind = Keyword.thy_decl | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1126 | in | 
| 426 | 1127 | Outer_Syntax.local_theory "foobar_trace" "traces a proposition" | 
| 1128 | kind (Parse.prop >> trace_prop) | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 1129 | end *} | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1130 | |
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1131 | text {*
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1132 |   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 | 1133 | see the proposition in the tracing buffer. | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1134 | *} | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1135 | |
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1136 | 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 | 1137 | |
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1138 | text {*
 | 
| 426 | 1139 |   Note that so far we used @{ML_ind thy_decl in Keyword} as the kind
 | 
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 1140 | 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 | 1141 | 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 | 1142 |   \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 | 1143 | 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 | 1144 | ``open up'' a proof in order to prove the proposition (for example | 
| 86 | 1145 |   \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 | 1146 |   \isacommand{function}). To achieve this kind of behaviour, you have to use
 | 
| 426 | 1147 |   the kind indicator @{ML_ind thy_goal in Keyword} and the function @{ML
 | 
| 1148 | "local_theory_to_proof" in Outer_Syntax} to set up the command. Note, | |
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 1149 |   however, once you change the ``kind'' of a command from @{ML thy_decl in
 | 
| 426 | 1150 |   Keyword} to @{ML thy_goal in Keyword} then the keyword file needs
 | 
| 219 
98d43270024f
more work on the simple inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
218diff
changeset | 1151 | to be re-created! | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1152 | |
| 327 
ce754ad78bc9
more work on the storing section
 Christian Urban <urbanc@in.tum.de> parents: 
326diff
changeset | 1153 |   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 | 1154 | 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 | 1155 |   it. Therefore in Line 9, we set the kind indicator to @{ML thy_goal in
 | 
| 426 | 1156 | Keyword}. | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1157 | *} | 
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1158 | |
| 114 
13fd0a83d3c3
properly handled linenumbers in ML-text and Isar-proofs
 Christian Urban <urbanc@in.tum.de> parents: 
108diff
changeset | 1159 | ML%linenosgray{*let
 | 
| 327 
ce754ad78bc9
more work on the storing section
 Christian Urban <urbanc@in.tum.de> parents: 
326diff
changeset | 1160 | 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 | 1161 | let | 
| 241 | 1162 | 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 | 1163 | in | 
| 422 | 1164 | Proof.theorem NONE (K I) [[(prop, [])]] lthy | 
| 327 
ce754ad78bc9
more work on the storing section
 Christian Urban <urbanc@in.tum.de> parents: 
326diff
changeset | 1165 | end | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1166 | |
| 426 | 1167 | val kind = Keyword.thy_goal | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1168 | in | 
| 426 | 1169 | Outer_Syntax.local_theory_to_proof "foobar_goal" "proves a proposition" | 
| 1170 | kind (Parse.prop >> goal_prop) | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 1171 | end *} | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1172 | |
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1173 | text {*
 | 
| 327 
ce754ad78bc9
more work on the storing section
 Christian Urban <urbanc@in.tum.de> parents: 
326diff
changeset | 1174 |   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 | 1175 | 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 | 1176 |   @{ML_ind read_prop in Syntax}, which converts a string into a proper proposition.
 | 
| 422 | 1177 |   In Line 6 the function @{ML_ind theorem in Proof} starts the proof for the
 | 
| 75 | 1178 |   proposition. Its argument @{ML NONE} stands for a locale (which we chose to
 | 
| 1179 |   omit); the argument @{ML "(K I)"} stands for a function that determines what
 | |
| 1180 | 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 | 1181 | 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 | 1182 | |
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1183 |   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 | 1184 | you obtain the following proof state | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1185 | *} | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1186 | |
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1187 | foobar_goal "True \<and> True" | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1188 | txt {*
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1189 |   \begin{minipage}{\textwidth}
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1190 |   @{subgoals [display]}
 | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1191 |   \end{minipage}\medskip
 | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
67diff
changeset | 1192 | |
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1193 | and can prove the proposition as follows. | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1194 | *} | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1195 | apply(rule conjI) | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1196 | apply(rule TrueI)+ | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1197 | done | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1198 | |
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1199 | text {*
 | 
| 327 
ce754ad78bc9
more work on the storing section
 Christian Urban <urbanc@in.tum.de> parents: 
326diff
changeset | 1200 |   {\bf TBD below}
 | 
| 74 | 1201 | |
| 394 | 1202 |   (FIXME: read a name and show how to store theorems; see @{ML_ind note in Local_Theory})
 | 
| 241 | 1203 | |
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1204 | *} | 
| 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 1205 | |
| 451 
fc074e669f9f
disabled foobar_prove; updated to new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
449diff
changeset | 1206 | ML {*
 | 
| 
fc074e669f9f
disabled foobar_prove; updated to new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
449diff
changeset | 1207 | structure Result = Proof_Data( | 
| 
fc074e669f9f
disabled foobar_prove; updated to new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
449diff
changeset | 1208 | type T = unit -> term | 
| 
fc074e669f9f
disabled foobar_prove; updated to new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
449diff
changeset | 1209 | fun init thy () = error "Result") | 
| 
fc074e669f9f
disabled foobar_prove; updated to new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
449diff
changeset | 1210 | |
| 
fc074e669f9f
disabled foobar_prove; updated to new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
449diff
changeset | 1211 | val result_cookie = (Result.get, Result.put, "Result.put") | 
| 
fc074e669f9f
disabled foobar_prove; updated to new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
449diff
changeset | 1212 | *} | 
| 
fc074e669f9f
disabled foobar_prove; updated to new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
449diff
changeset | 1213 | |
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1214 | ML{*let
 | 
| 324 
4172c0743cf2
updated foobar_proof example
 Christian Urban <urbanc@in.tum.de> parents: 
322diff
changeset | 1215 | fun after_qed thm_name thms lthy = | 
| 394 | 1216 | Local_Theory.note (thm_name, (flat thms)) lthy |> snd | 
| 324 
4172c0743cf2
updated foobar_proof example
 Christian Urban <urbanc@in.tum.de> parents: 
322diff
changeset | 1217 | |
| 
4172c0743cf2
updated foobar_proof example
 Christian Urban <urbanc@in.tum.de> parents: 
322diff
changeset | 1218 | fun setup_proof (thm_name, (txt, pos)) lthy = | 
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1219 | let | 
| 451 
fc074e669f9f
disabled foobar_prove; updated to new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
449diff
changeset | 1220 |      val trm = Code_Runtime.value lthy result_cookie ("", txt)
 | 
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1221 | in | 
| 422 | 1222 | Proof.theorem NONE (after_qed thm_name) [[(trm,[])]] lthy | 
| 324 
4172c0743cf2
updated foobar_proof example
 Christian Urban <urbanc@in.tum.de> parents: 
322diff
changeset | 1223 | end | 
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1224 | |
| 426 | 1225 | val parser = Parse_Spec.opt_thm_name ":" -- Parse.ML_source | 
| 451 
fc074e669f9f
disabled foobar_prove; updated to new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
449diff
changeset | 1226 | |
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1227 | in | 
| 426 | 1228 | Outer_Syntax.local_theory_to_proof "foobar_prove" "proving a proposition" | 
| 1229 | Keyword.thy_goal (parser >> setup_proof) | |
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1230 | end*} | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1231 | |
| 451 
fc074e669f9f
disabled foobar_prove; updated to new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
449diff
changeset | 1232 | (* | 
| 324 
4172c0743cf2
updated foobar_proof example
 Christian Urban <urbanc@in.tum.de> parents: 
322diff
changeset | 1233 | foobar_prove test: {* @{prop "True"} *}
 | 
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1234 | apply(rule TrueI) | 
| 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1235 | done | 
| 451 
fc074e669f9f
disabled foobar_prove; updated to new Isabelle
 Christian Urban <urbanc@in.tum.de> parents: 
449diff
changeset | 1236 | *) | 
| 321 
e450fa467e3f
polished the commands section
 Christian Urban <urbanc@in.tum.de> parents: 
319diff
changeset | 1237 | |
| 322 | 1238 | (* | 
| 1239 | ML {*
 | |
| 1240 | structure TacticData = ProofDataFun | |
| 1241 | ( | |
| 1242 | type T = thm list -> tactic; | |
| 1243 | fun init _ = undefined; | |
| 366 | 1244 | ) | 
| 322 | 1245 | |
| 1246 | val set_tactic = TacticData.put; | |
| 1247 | *} | |
| 1248 | ||
| 1249 | ML {*
 | |
| 1250 |   TacticData.get @{context}
 | |
| 1251 | *} | |
| 1252 | ||
| 1253 | ML {* Method.set_tactic  *}
 | |
| 1254 | ML {* fun tactic (facts: thm list) : tactic = (atac 1) *}
 | |
| 1255 | ML {* Context.map_proof *}
 | |
| 1256 | ML {* ML_Context.expression *}
 | |
| 1257 | ML {* METHOD *}
 | |
| 1258 | ||
| 1259 | ||
| 1260 | ML {* 
 | |
| 1261 | fun myexpression pos bind body txt = | |
| 1262 | let | |
| 1263 |   val _ = tracing ("bind)" ^ bind)
 | |
| 1264 |   val _ = tracing ("body)" ^ body)
 | |
| 1265 |   val _ = tracing ("txt)"  ^ txt)
 | |
| 1266 |   val _ = tracing ("result) " ^ "Context.set_thread_data (SOME (let " ^ bind ^ " = " ^ txt ^ " in " ^ body ^
 | |
| 1267 | " end (ML_Context.the_generic_context ())));") | |
| 1268 | in | |
| 1269 | ML_Context.exec (fn () => ML_Context.eval false pos | |
| 1270 |     ("Context.set_thread_data (SOME (let " ^ bind ^ " = " ^ txt ^ " in " ^ body ^
 | |
| 1271 | " end (ML_Context.the_generic_context ())));")) | |
| 1272 | end | |
| 1273 | *} | |
| 319 
6bce4acf7f2a
added file for producing a keyword file
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 1274 | |
| 
6bce4acf7f2a
added file for producing a keyword file
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 1275 | |
| 322 | 1276 | ML {*
 | 
| 1277 | fun ml_tactic (txt, pos) ctxt = | |
| 1278 | let | |
| 1279 | val ctxt' = ctxt |> Context.proof_map | |
| 1280 | (myexpression pos | |
| 1281 | "fun tactic (facts: thm list) : tactic" | |
| 1282 | "Context.map_proof (Method.set_tactic tactic)" txt); | |
| 1283 | in | |
| 1284 | Context.setmp_thread_data (SOME (Context.Proof ctxt)) (TacticData.get ctxt') | |
| 1285 | end; | |
| 1286 | *} | |
| 1287 | ||
| 1288 | ML {*
 | |
| 1289 | fun tactic3 (txt, pos) ctxt = | |
| 1290 | let | |
| 1291 |     val _ = tracing ("1) " ^ txt )
 | |
| 1292 | in | |
| 1293 | METHOD (ml_tactic (txt, pos) ctxt; K (atac 1)) | |
| 1294 | end | |
| 1295 | *} | |
| 1296 | ||
| 1297 | setup {*
 | |
| 426 | 1298 | Method.setup (Binding.name "tactic3") (Scan.lift (Parse.position Args.name) | 
| 322 | 1299 | >> tactic3) | 
| 1300 | "ML tactic as proof method" | |
| 1301 | *} | |
| 1302 | ||
| 1303 | lemma "A \<Longrightarrow> A" | |
| 1304 | apply(tactic3 {* (atac 1)  *})
 | |
| 1305 | done | |
| 1306 | ||
| 1307 | ML {* 
 | |
| 1308 | (ML_Context.the_generic_context ()) | |
| 1309 | *} | |
| 1310 | ||
| 1311 | ML {*
 | |
| 1312 | Context.set_thread_data; | |
| 1313 | ML_Context.the_generic_context | |
| 1314 | *} | |
| 1315 | ||
| 1316 | lemma "A \<Longrightarrow> A" | |
| 1317 | ML_prf {*
 | |
| 1318 | 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 ()))); | |
| 1319 | *} | |
| 1320 | ||
| 1321 | ML {*
 | |
| 1322 | Context.set_thread_data (SOME ((let fun tactic (facts: thm list) : tactic = (atac 1) in 3 end) (ML_Context.the_generic_context ()))); | |
| 1323 | *} | |
| 1324 | ||
| 1325 | ML {*
 | |
| 1326 | Context.set_thread_data (SOME (let | |
| 1327 | fun tactic (facts: thm list) : tactic = (atac 1) | |
| 1328 | in | |
| 1329 | Context.map_proof (Method.set_tactic tactic) | |
| 1330 | end | |
| 1331 | (ML_Context.the_generic_context ()))); | |
| 1332 | *} | |
| 1333 | ||
| 1334 | ||
| 1335 | ML {*
 | |
| 1336 | let | |
| 1337 | fun tactic (facts: thm list) : tactic = atac | |
| 1338 | in | |
| 1339 | Context.map_proof (Method.set_tactic tactic) | |
| 1340 | end *} | |
| 1341 | ||
| 1342 | end *} | |
| 1343 | ||
| 1344 | ML {* Toplevel.program (fn () => 
 | |
| 1345 | (ML_Context.expression Position.none "val plus : int" "3 + 4" "1" (Context.Proof @{context})))*}
 | |
| 1346 | ||
| 1347 | ||
| 1348 | ML {*
 | |
| 1349 | fun ml_tactic (txt, pos) ctxt = | |
| 1350 | let | |
| 1351 | val ctxt' = ctxt |> Context.proof_map | |
| 1352 | (ML_Context.expression pos | |
| 1353 | "fun tactic (facts: thm list) : tactic" | |
| 1354 | "Context.map_proof (Method.set_tactic tactic)" txt); | |
| 1355 | in Context.setmp_thread_data (SOME (Context.Proof ctxt)) (TacticData.get ctxt') end; | |
| 1356 | ||
| 1357 | *} | |
| 1358 | ||
| 1359 | ML {*
 | |
| 1360 | 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 ()))); | |
| 1361 | *} | |
| 1362 | *) | |
| 319 
6bce4acf7f2a
added file for producing a keyword file
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 1363 | |
| 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 | 1364 | 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 | 1365 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1366 | text {*
 | 
| 207 | 1367 | (FIXME: maybe move to after the tactic section) | 
| 1368 | ||
| 221 | 1369 | 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 | 1370 |   in \isacommand{apply}. To print out all currently known methods you can use the 
 | 
| 192 | 1371 | 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 | 1372 | |
| 207 | 1373 |   \begin{isabelle}
 | 
| 1374 |   \isacommand{print\_methods}\\
 | |
| 1375 |   @{text "> methods:"}\\
 | |
| 1376 |   @{text ">   -:  do nothing (insert current facts only)"}\\
 | |
| 1377 |   @{text ">   HOL.default:  apply some intro/elim rule (potentially classical)"}\\
 | |
| 1378 |   @{text ">   ..."}
 | |
| 1379 |   \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 | 1380 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1381 | 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 | 1382 | *} | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1383 | |
| 244 
dc95a56b1953
fixed the problem with double definition of even and odd
 Christian Urban <urbanc@in.tum.de> parents: 
241diff
changeset | 1384 | method_setup %gray foo = | 
| 181 
5baaabe1ab92
updated to new method_setup
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 1385 |  {* Scan.succeed
 | 
| 
5baaabe1ab92
updated to new method_setup
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 1386 |       (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 | 1387 | "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 | 1388 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1389 | text {*
 | 
| 286 | 1390 |   It defines the method @{text foo}, which takes no arguments (therefore the
 | 
| 207 | 1391 |   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 | 1392 |   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 | 1393 |   @{ML_ind SIMPLE_METHOD in Method}
 | 
| 287 | 1394 |   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 | 1395 | *} | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1396 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1397 | 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 | 1398 | 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 | 1399 | 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 | 1400 | 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 | 1401 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1402 |   \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 | 1403 |   @{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 | 1404 |   \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 | 1405 | (*<*)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 | 1406 | |
| 421 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1407 | method_setup test = {* Scan.lift (Scan.succeed (K Method.succeed)) *} {* bla *}
 | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1408 | |
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1409 | lemma "True" | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1410 | apply(test) | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1411 | oops | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1412 | |
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1413 | method_setup joker = {* Scan.lift (Scan.succeed (fn ctxt => Method.cheating true ctxt)) *} {* bla *}
 | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1414 | |
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1415 | lemma "False" | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1416 | apply(joker) | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1417 | oops | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1418 | |
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1419 | text {* if true is set then always works *}
 | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1420 | |
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1421 | ML {* atac *} 
 | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1422 | |
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1423 | method_setup first_atac = {* Scan.lift (Scan.succeed (K (SIMPLE_METHOD (atac 1)))) *} {* bla *}
 | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1424 | |
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1425 | ML {* HEADGOAL *}
 | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1426 | |
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1427 | lemma "A \<Longrightarrow> A" | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1428 | apply(first_atac) | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1429 | oops | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1430 | |
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1431 | method_setup my_atac = {* Scan.lift (Scan.succeed (K (SIMPLE_METHOD' atac))) *} {* bla *}
 | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1432 | |
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1433 | lemma "A \<Longrightarrow> A" | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1434 | apply(my_atac) | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1435 | oops | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1436 | |
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1437 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1438 | |
| 319 
6bce4acf7f2a
added file for producing a keyword file
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 1439 | |
| 
6bce4acf7f2a
added file for producing a keyword file
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 1440 | |
| 
6bce4acf7f2a
added file for producing a keyword file
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 1441 | |
| 
6bce4acf7f2a
added file for producing a keyword file
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 1442 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1443 | (* | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1444 | ML {* SIMPLE_METHOD *}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1445 | ML {* METHOD *}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1446 | 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 | 1447 | ML {* Scan.succeed  *}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1448 | *) | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1449 | |
| 421 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1450 | ML {* resolve_tac *}
 | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1451 | |
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1452 | method_setup myrule = | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1453 |   {* Scan.lift (Scan.succeed (K (METHOD (fn thms => resolve_tac thms 1)))) *}
 | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1454 |   {* bla *}
 | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1455 | |
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1456 | lemma | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1457 | assumes a: "A \<Longrightarrow> B \<Longrightarrow> C" | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1458 | shows "C" | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1459 | using a | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1460 | apply(myrule) | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1461 | oops | 
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1462 | |
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1463 | |
| 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1464 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
183diff
changeset | 1465 | text {*
 | 
| 421 
620a24bf954a
added a section to the introduction; described @{make_string}
 Christian Urban <urbanc@in.tum.de> parents: 
414diff
changeset | 1466 | (********************************************************) | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
183diff
changeset | 1467 | (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 | 1468 | *} | 
| 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 | 1469 | |
| 220 | 1470 | end |