ProgTutorial/Parsing.thy
author Christian Urban <urbanc@in.tum.de>
Sat, 21 Mar 2009 12:35:03 +0100
changeset 192 2fff636e1fa0
parent 189 069d525f8f1d
child 193 ffd93dcc269d
permissions -rw-r--r--
some polishing
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Parsing
121
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
     2
imports Base "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
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
     5
4
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
chapter {* Parsing *}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
text {*
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    10
  Isabelle distinguishes between \emph{outer} and \emph{inner} syntax. 
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    11
  Theory commands, such as \isacommand{definition}, \isacommand{inductive} and so
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    12
  on, belong to the outer syntax, whereas items inside double quotation marks, such 
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    13
  as terms, types and so on, belong to the inner syntax. For parsing inner syntax, 
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 149
diff changeset
    14
  Isabelle uses a rather general and sophisticated algorithm, which 
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    15
  is driven by priority grammars. Parsers for outer syntax are built up by functional
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    16
  parsing combinators. These combinators are a well-established technique for parsing, 
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
    17
  which has, for example, been described in Paulson's classic ML-book \cite{paulson-ml2}.
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
    18
  Isabelle developers are usually concerned with writing these outer syntax parsers, 
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
    19
  either for new definitional packages or for calling tactics with specific arguments. 
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
    20
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
    21
  \begin{readmore}
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
    22
  The library 
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
    23
  for writing parser combinators is split up, roughly, into two parts. 
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    24
  The first part consists of a collection of generic parser combinators defined
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    25
  in the structure @{ML_struct Scan} in the file 
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    26
  @{ML_file "Pure/General/scan.ML"}. The second part of the library consists of 
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    27
  combinators for dealing with specific token types, which are defined in the 
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    28
  structure @{ML_struct OuterParse} in the file 
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    29
  @{ML_file "Pure/Isar/outer_parse.ML"}.
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
    30
  \end{readmore}
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    31
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    32
*}
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    33
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
    34
section {* Building Generic Parsers *}
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    35
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    36
text {*
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    37
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    38
  Let us first have a look at parsing strings using generic parsing combinators. 
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
    39
  The function @{ML "$$"} takes a string as argument and will ``consume'' this string from 
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    40
  a given input list of strings. ``Consume'' in this context means that it will 
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    41
  return a pair consisting of this string and the rest of the input list. 
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    42
  For example:
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    43
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
    44
  @{ML_response [display,gray] "($$ \"h\") (explode \"hello\")" "(\"h\", [\"e\", \"l\", \"l\", \"o\"])"}
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
    45
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
    46
  @{ML_response [display,gray] "($$ \"w\") (explode \"world\")" "(\"w\", [\"o\", \"r\", \"l\", \"d\"])"}
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    47
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
    48
  The function @{ML "$$"} will either succeed (as in the two examples above) or raise the exception 
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
    49
  @{text "FAIL"} if no string can be consumed. For example trying to parse
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    50
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
    51
  @{ML_response_fake [display,gray] "($$ \"x\") (explode \"world\")" 
41
b11653b11bd3 further progress on the parsing section and tuning on the antiqu's
Christian Urban <urbanc@in.tum.de>
parents: 40
diff changeset
    52
                               "Exception FAIL raised"}
b11653b11bd3 further progress on the parsing section and tuning on the antiqu's
Christian Urban <urbanc@in.tum.de>
parents: 40
diff changeset
    53
  
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
    54
  will raise the exception @{text "FAIL"}.
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    55
  There are three exceptions used in the parsing combinators:
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    56
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    57
  \begin{itemize}
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
    58
  \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: 40
diff changeset
    59
  might be explored. 
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
    60
  \item @{text "MORE"} indicates that there is not enough input for the parser. For example 
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
    61
  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: 59
diff changeset
    62
  \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: 105
diff changeset
    63
  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: 38
diff changeset
    64
  \end{itemize}
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    65
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
    66
  However, note that these exceptions are private to the parser and cannot be accessed
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
    67
  by the programmer (for example to handle them).
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
    68
  
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
    69
  Slightly more general than the parser @{ML "$$"} is the function @{ML
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
    70
  Scan.one}, in that it takes a predicate as argument and then parses exactly
52
Christian Urban <urbanc@in.tum.de>
parents: 50
diff changeset
    71
  one item from the input list satisfying this predicate. For example the
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
    72
  following parser either consumes an @{text [quotes] "h"} or a @{text
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
    73
  [quotes] "w"}:
41
b11653b11bd3 further progress on the parsing section and tuning on the antiqu's
Christian Urban <urbanc@in.tum.de>
parents: 40
diff changeset
    74
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    75
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
    76
@{ML_response [display,gray] 
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
    77
"let 
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
    78
  val hw = Scan.one (fn x => x = \"h\" orelse x = \"w\")
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    79
  val input1 = (explode \"hello\")
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    80
  val input2 = (explode \"world\")
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    81
in
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    82
    (hw input1, hw input2)
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    83
end"
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    84
    "((\"h\", [\"e\", \"l\", \"l\", \"o\"]),(\"w\", [\"o\", \"r\", \"l\", \"d\"]))"}
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    85
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
    86
  Two parser can be connected in sequence by using the function @{ML "--"}. 
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
    87
  For example parsing @{text "h"}, @{text "e"} and @{text "l"} in this 
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
    88
  sequence you can achieve by:
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    89
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
    90
  @{ML_response [display,gray] "(($$ \"h\") -- ($$ \"e\") -- ($$ \"l\")) (explode \"hello\")"
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    91
                          "(((\"h\", \"e\"), \"l\"), [\"l\", \"o\"])"}
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    92
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    93
  Note how the result of consumed strings builds up on the left as nested pairs.  
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    94
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
    95
  If, as in the previous example, you want to parse a particular string, 
128
693711a0c702 polished
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
    96
  then you should use the function @{ML Scan.this_string}:
56
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    97
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
    98
  @{ML_response [display,gray] "Scan.this_string \"hell\" (explode \"hello\")"
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
    99
                          "(\"hell\", [\"o\"])"}
56
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   100
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   101
  Parsers that explore alternatives can be constructed using the function @{ML
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   102
  "||"}. For example, the parser @{ML "(p || q)" for p q} returns the
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
   103
  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: 101
diff changeset
   104
  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: 54
diff changeset
   105
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   106
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   107
@{ML_response [display,gray] 
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   108
"let 
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   109
  val hw = ($$ \"h\") || ($$ \"w\")
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   110
  val input1 = (explode \"hello\")
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   111
  val input2 = (explode \"world\")
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   112
in
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   113
  (hw input1, hw input2)
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   114
end"
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   115
  "((\"h\", [\"e\", \"l\", \"l\", \"o\"]), (\"w\", [\"o\", \"r\", \"l\", \"d\"]))"}
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   116
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   117
  The functions @{ML "|--"} and @{ML "--|"} work like the sequencing function 
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   118
  for parsers, except that they discard the item being parsed by the first (respectively second)
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   119
  parser. For example:
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   120
  
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   121
@{ML_response [display,gray]
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   122
"let 
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
   123
  val just_e = ($$ \"h\") |-- ($$ \"e\") 
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
   124
  val just_h = ($$ \"h\") --| ($$ \"e\") 
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   125
  val input = (explode \"hello\")  
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   126
in 
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
   127
  (just_e input, just_h input)
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   128
end"
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   129
  "((\"e\", [\"l\", \"l\", \"o\"]),(\"h\", [\"l\", \"l\", \"o\"]))"}
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   130
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   131
  The parser @{ML "Scan.optional p x" for p x} returns the result of the parser 
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
   132
  @{text "p"}, if it succeeds; otherwise it returns 
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   133
  the default value @{text "x"}. For example:
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   134
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   135
@{ML_response [display,gray]
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   136
"let 
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   137
  val p = Scan.optional ($$ \"h\") \"x\"
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   138
  val input1 = (explode \"hello\")
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   139
  val input2 = (explode \"world\")  
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   140
in 
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   141
  (p input1, p input2)
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   142
end" 
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   143
 "((\"h\", [\"e\", \"l\", \"l\", \"o\"]), (\"x\", [\"w\", \"o\", \"r\", \"l\", \"d\"]))"}
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   144
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   145
  The function @{ML Scan.option} works similarly, except no default value can
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   146
  be given. Instead, the result is wrapped as an @{text "option"}-type. For example:
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   147
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   148
@{ML_response [display,gray]
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   149
"let 
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   150
  val p = Scan.option ($$ \"h\")
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   151
  val input1 = (explode \"hello\")
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   152
  val input2 = (explode \"world\")  
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   153
in 
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   154
  (p input1, p input2)
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   155
end" "((SOME \"h\", [\"e\", \"l\", \"l\", \"o\"]), (NONE, [\"w\", \"o\", \"r\", \"l\", \"d\"]))"}
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   156
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   157
  The function @{ML "!!"} helps to produce appropriate error messages
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   158
  during parsing. For example if you want to parse that @{text p} is immediately 
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
   159
  followed by @{text q}, or start a completely different parser @{text r},
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   160
  you might write:
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   161
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   162
  @{ML [display,gray] "(p -- q) || r" for p q r}
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   163
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   164
  However, this parser is problematic for producing an appropriate error
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   165
  message, in case the parsing of @{ML "(p -- q)" for p q} fails. Because in
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   166
  that case you lose the information that @{text p} should be followed by
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   167
  @{text q}. To see this consider the case in which @{text p} is present in
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   168
  the input, but not @{text q}. That means @{ML "(p -- q)" for p q} will fail
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   169
  and the alternative parser @{text r} will be tried. However in many
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   170
  circumstance this will be the wrong parser for the input ``p-followed-by-q''
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   171
  and therefore will also fail. The error message is then caused by the
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   172
  failure of @{text r}, not by the absence of @{text q} in the input. This
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   173
  kind of situation can be avoided when using the function @{ML "!!"}. 
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   174
  This function aborts the whole process of parsing in case of a
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   175
  failure and prints an error message. For example if you invoke the parser
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   176
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   177
  
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   178
  @{ML [display,gray] "(!! (fn _ => \"foo\") ($$ \"h\"))"}
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   179
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
   180
  on @{text [quotes] "hello"}, the parsing succeeds
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   181
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   182
  @{ML_response [display,gray] 
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   183
                "(!! (fn _ => \"foo\") ($$ \"h\")) (explode \"hello\")" 
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   184
                "(\"h\", [\"e\", \"l\", \"l\", \"o\"])"}
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   185
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   186
  but if you invoke it on @{text [quotes] "world"}
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   187
  
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   188
  @{ML_response_fake [display,gray] "(!! (fn _ => \"foo\") ($$ \"h\")) (explode \"world\")"
41
b11653b11bd3 further progress on the parsing section and tuning on the antiqu's
Christian Urban <urbanc@in.tum.de>
parents: 40
diff changeset
   189
                               "Exception ABORT raised"}
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   190
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   191
  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: 116
diff changeset
   192
  see the error message properly, you need to prefix the parser with the function 
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   193
  @{ML "Scan.error"}. For example:
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   194
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   195
  @{ML_response_fake [display,gray] "Scan.error (!! (fn _ => \"foo\") ($$ \"h\"))"
41
b11653b11bd3 further progress on the parsing section and tuning on the antiqu's
Christian Urban <urbanc@in.tum.de>
parents: 40
diff changeset
   196
                               "Exception Error \"foo\" raised"}
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   197
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   198
  This ``prefixing'' is usually done by wrappers such as @{ML "OuterSyntax.command"} 
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   199
  (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: 39
diff changeset
   200
  
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   201
  Let us now return to our example of parsing @{ML "(p -- q) || r" for p q
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   202
  r}. If you want to generate the correct error message for p-followed-by-q,
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   203
  then you have to write:
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   204
*}
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   205
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   206
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: 132
diff changeset
   207
let 
3e94ccc0f31e polishing and start of the section about attributes
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   208
  val err_msg = (fn _ => p ^ " is not followed by " ^ q)
3e94ccc0f31e polishing and start of the section about attributes
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   209
in
3e94ccc0f31e polishing and start of the section about attributes
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   210
  ($$ p -- (!! err_msg ($$ q))) || ($$ r -- $$ r)
3e94ccc0f31e polishing and start of the section about attributes
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   211
end *}
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   212
41
b11653b11bd3 further progress on the parsing section and tuning on the antiqu's
Christian Urban <urbanc@in.tum.de>
parents: 40
diff changeset
   213
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   214
text {*
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   215
  Running this parser with the @{text [quotes] "h"} and @{text [quotes] "e"}, and 
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   216
  the input @{text [quotes] "holle"} 
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   217
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   218
  @{ML_response_fake [display,gray] "Scan.error (p_followed_by_q \"h\" \"e\" \"w\") (explode \"holle\")"
41
b11653b11bd3 further progress on the parsing section and tuning on the antiqu's
Christian Urban <urbanc@in.tum.de>
parents: 40
diff changeset
   219
                               "Exception ERROR \"h is not followed by e\" raised"} 
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   220
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   221
  produces the correct error message. Running it with
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   222
 
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   223
  @{ML_response [display,gray] "Scan.error (p_followed_by_q \"h\" \"e\" \"w\") (explode \"wworld\")"
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   224
                          "((\"w\", \"w\"), [\"o\", \"r\", \"l\", \"d\"])"}
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   225
  
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   226
  yields the expected parsing. 
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   227
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
   228
  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: 101
diff changeset
   229
  often as it succeeds. For example:
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   230
  
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   231
  @{ML_response [display,gray] "Scan.repeat ($$ \"h\") (explode \"hhhhello\")" 
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   232
                "([\"h\", \"h\", \"h\", \"h\"], [\"e\", \"l\", \"l\", \"o\"])"}
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   233
  
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   234
  Note that @{ML "Scan.repeat"} stores the parsed items in a list. The function
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
   235
  @{ML "Scan.repeat1"} 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: 40
diff changeset
   236
  succeeds at least once.
48
609f9ef73494 fixed FIXME's in fake responses
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
   237
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
   238
  Also note that the parser would have aborted with the exception @{text MORE}, if
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   239
  you had run it only on just @{text [quotes] "hhhh"}. This can be avoided by using
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   240
  the wrapper @{ML Scan.finite} and the ``stopper-token'' @{ML Symbol.stopper}. With
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   241
  them you can write:
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   242
  
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   243
  @{ML_response [display,gray] "Scan.finite Symbol.stopper (Scan.repeat ($$ \"h\")) (explode \"hhhh\")" 
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   244
                "([\"h\", \"h\", \"h\", \"h\"], [])"}
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   245
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   246
  @{ML Symbol.stopper} is the ``end-of-input'' indicator for parsing strings;
128
693711a0c702 polished
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   247
  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: 60
diff changeset
   248
  manually wrapping is often already done by the surrounding infrastructure. 
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   249
56
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   250
  The function @{ML Scan.repeat} can be used with @{ML Scan.one} to read any 
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   251
  string as in
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   252
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   253
  @{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: 54
diff changeset
   254
"let 
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   255
   val p = Scan.repeat (Scan.one Symbol.not_eof)
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   256
   val input = (explode \"foo bar foo\") 
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   257
in
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   258
   Scan.finite Symbol.stopper p input
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   259
end" 
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   260
"([\"f\", \"o\", \"o\", \" \", \"b\", \"a\", \"r\", \" \", \"f\", \"o\", \"o\"], [])"}
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   261
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   262
  where the function @{ML Symbol.not_eof} 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: 60
diff changeset
   263
  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: 54
diff changeset
   264
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   265
  The function @{ML "Scan.unless p q" for p q} 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: 59
diff changeset
   266
  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: 54
diff changeset
   267
  
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   268
  @{ML_response_fake_both [display,gray] "Scan.unless ($$ \"h\") ($$ \"w\") (explode \"hello\")"
56
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   269
                               "Exception FAIL raised"}
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   270
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   271
  fails, while
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   272
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   273
  @{ML_response [display,gray] "Scan.unless ($$ \"h\") ($$ \"w\") (explode \"world\")"
56
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   274
                          "(\"w\",[\"o\", \"r\", \"l\", \"d\"])"}
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   275
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   276
  succeeds. 
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   277
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   278
  The functions @{ML Scan.repeat} and @{ML Scan.unless} can be combined to read any
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   279
  input until a certain marker symbol is reached. In the example below the marker
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: 59
diff changeset
   280
  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: 54
diff changeset
   281
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   282
  @{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: 54
diff changeset
   283
"let 
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   284
  val p = Scan.repeat (Scan.unless ($$ \"*\") (Scan.one Symbol.not_eof))
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   285
  val input1 = (explode \"fooooo\")
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   286
  val input2 = (explode \"foo*ooo\")
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   287
in
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   288
  (Scan.finite Symbol.stopper p input1, 
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   289
   Scan.finite Symbol.stopper p input2)
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   290
end"
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   291
"(([\"f\", \"o\", \"o\", \"o\", \"o\", \"o\"], []),
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   292
 ([\"f\", \"o\", \"o\"], [\"*\", \"o\", \"o\", \"o\"]))"}
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   293
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   294
  After parsing is done, you nearly always want to apply a function on the parsed 
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   295
  items. One way to do this is the function @{ML "(p >> f)" for p f}, which runs 
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
   296
  first the parser @{text p} and upon successful completion applies the 
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
   297
  function @{text f} to the result. For example
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   298
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   299
@{ML_response [display,gray]
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   300
"let 
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   301
  fun double (x,y) = (x ^ x, y ^ y) 
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   302
in
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   303
  (($$ \"h\") -- ($$ \"e\") >> double) (explode \"hello\")
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   304
end"
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   305
"((\"hh\", \"ee\"), [\"l\", \"l\", \"o\"])"}
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   306
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   307
  doubles the two parsed input strings; or
59
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   308
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   309
  @{ML_response [display,gray] 
59
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   310
"let 
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   311
   val p = Scan.repeat (Scan.one Symbol.not_eof)
59
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   312
   val input = (explode \"foo bar foo\") 
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   313
in
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   314
   Scan.finite Symbol.stopper (p >> implode) input
59
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   315
end" 
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   316
"(\"foo bar foo\",[])"}
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   317
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: 59
diff changeset
   318
  where the single-character strings in the parsed output are transformed
59
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   319
  back into one string.
41
b11653b11bd3 further progress on the parsing section and tuning on the antiqu's
Christian Urban <urbanc@in.tum.de>
parents: 40
diff changeset
   320
 
125
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   321
  The function @{ML Scan.ahead} parses some input, but leaves the original
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   322
  input unchanged. For example:
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   323
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   324
  @{ML_response [display,gray]
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   325
  "Scan.ahead (Scan.this_string \"foo\") (explode \"foo\")" 
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   326
  "(\"foo\", [\"f\", \"o\", \"o\"])"} 
56
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   327
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   328
  The function @{ML Scan.lift} takes a parser and a pair as arguments. This function applies
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   329
  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: 39
diff changeset
   330
  untouched. For example
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   331
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   332
@{ML_response [display,gray]
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   333
"Scan.lift (($$ \"h\") -- ($$ \"e\")) (1,(explode \"hello\"))"
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   334
"((\"h\", \"e\"), (1, [\"l\", \"l\", \"o\"]))"}
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   335
43
02f76f1b6e7b added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   336
  (FIXME: In which situations is this useful? Give examples.) 
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
   337
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
   338
  \begin{exercise}\label{ex:scancmts}
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
   339
  Write a parser that parses an input string so that any comment enclosed
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
   340
  inside @{text "(*\<dots>*)"} is replaced by a the same comment but enclosed inside
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
   341
  @{text "(**\<dots>**)"} in the output string. To enclose a string, you can use the
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
   342
  function @{ML "enclose s1 s2 s" for s1 s2 s} which produces the string @{ML
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
   343
  "s1 ^ s ^ s2" for s1 s2 s}.
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
   344
  \end{exercise}
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   345
*}
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   346
41
b11653b11bd3 further progress on the parsing section and tuning on the antiqu's
Christian Urban <urbanc@in.tum.de>
parents: 40
diff changeset
   347
section {* Parsing Theory Syntax *}
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   348
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   349
text {*
188
8939b8fd8603 added lots of FIXMES
Christian Urban <urbanc@in.tum.de>
parents: 186
diff changeset
   350
  (FIXME: context parser)
8939b8fd8603 added lots of FIXMES
Christian Urban <urbanc@in.tum.de>
parents: 186
diff changeset
   351
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   352
  Most of the time, however, Isabelle developers have to deal with parsing
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 149
diff changeset
   353
  tokens, not strings.  These token parsers have the type:
128
693711a0c702 polished
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   354
*}
693711a0c702 polished
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   355
  
693711a0c702 polished
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   356
ML{*type 'a parser = OuterLex.token list -> 'a * OuterLex.token list*}
693711a0c702 polished
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   357
693711a0c702 polished
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   358
text {*
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
   359
  The reason for using token parsers is that theory syntax, as well as the
128
693711a0c702 polished
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   360
  parsers for the arguments of proof methods, use the type @{ML_type
693711a0c702 polished
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   361
  OuterLex.token} (which is identical to the type @{ML_type
693711a0c702 polished
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   362
  OuterParse.token}).  However, there are also handy parsers for
693711a0c702 polished
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   363
  ML-expressions and ML-files.
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   364
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   365
  \begin{readmore}
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   366
  The parser functions for the theory syntax are contained in the structure
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   367
  @{ML_struct OuterParse} defined in the file @{ML_file  "Pure/Isar/outer_parse.ML"}.
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   368
  The definition for tokens is in the file @{ML_file "Pure/Isar/outer_lex.ML"}.
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   369
  \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: 43
diff changeset
   370
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   371
  The structure @{ML_struct OuterLex} defines several kinds of tokens (for example 
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   372
  @{ML "Ident" in OuterLex} for identifiers, @{ML "Keyword" in OuterLex} for keywords and
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   373
  @{ML "Command" in OuterLex} for commands). Some token parsers take into account the 
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   374
  kind of tokens.
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   375
*}  
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   376
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   377
text {*
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   378
  The first example shows how to generate a token list out of a string using
128
693711a0c702 polished
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   379
  the function @{ML "OuterSyntax.scan"}. It is given the argument @{ML "Position.none"}
693711a0c702 polished
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   380
  since, at the moment, we are not interested in generating
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   381
  precise error messages. The following code
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: 43
diff changeset
   382
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   383
@{ML_response_fake [display,gray] "OuterSyntax.scan Position.none \"hello world\"" 
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   384
"[Token (\<dots>,(Ident, \"hello\"),\<dots>), 
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   385
 Token (\<dots>,(Space, \" \"),\<dots>), 
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   386
 Token (\<dots>,(Ident, \"world\"),\<dots>)]"}
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   387
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   388
  produces three tokens where the first and the last are identifiers, since
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
   389
  @{text [quotes] "hello"} and @{text [quotes] "world"} do not match any
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   390
  other syntactic category.\footnote{Note that because of a possible a bug in
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   391
  the PolyML runtime system the result is printed as @{text [quotes] "?"}, instead of
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   392
  the tokens.} The second indicates a space.
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   393
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   394
  Many parsing functions later on will require spaces, comments and the like
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   395
  to have already been filtered out.  So from now on we are going to use 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: 101
diff changeset
   396
  functions @{ML filter} and @{ML OuterLex.is_proper} do this. 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: 43
diff changeset
   397
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   398
@{ML_response_fake [display,gray]
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   399
"let
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   400
   val input = OuterSyntax.scan Position.none \"hello world\"
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   401
in
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   402
   filter OuterLex.is_proper input
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   403
end" 
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   404
"[Token (\<dots>,(Ident, \"hello\"), \<dots>), Token (\<dots>,(Ident, \"world\"), \<dots>)]"}
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   405
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   406
  For convenience we define the function:
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   407
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   408
*}
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   409
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   410
ML{*fun filtered_input str = 
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
   411
  filter OuterLex.is_proper (OuterSyntax.scan Position.none str) *}
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   412
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   413
text {*
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   414
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   415
  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: 43
diff changeset
   416
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   417
@{ML_response_fake [display,gray] 
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   418
"filtered_input \"inductive | for\"" 
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   419
"[Token (\<dots>,(Command, \"inductive\"),\<dots>), 
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   420
 Token (\<dots>,(Keyword, \"|\"),\<dots>), 
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   421
 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: 43
diff changeset
   422
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   423
  you obtain a list consisting of only a command and two keyword tokens.
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   424
  If you want to see which keywords and commands are currently known to Isabelle, type in
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   425
  the following code (you might have to adjust the @{ML print_depth} in order to
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
   426
  see the complete list):
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
   427
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   428
@{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: 44
diff changeset
   429
"let 
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
   430
  val (keywords, commands) = OuterKeyword.get_lexicons ()
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
   431
in 
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
   432
  (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: 44
diff changeset
   433
end" 
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   434
"([\"}\", \"{\", \<dots>], [\"\<rightleftharpoons>\", \"\<leftharpoondown>\", \<dots>])"}
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   435
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   436
  The parser @{ML "OuterParse.$$$"} parses a single keyword. For example:
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   437
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   438
@{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: 43
diff changeset
   439
"let 
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   440
  val input1 = filtered_input \"where for\"
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   441
  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: 43
diff changeset
   442
in 
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents: 43
diff changeset
   443
  (OuterParse.$$$ \"where\" input1, OuterParse.$$$ \"|\" input2)
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents: 43
diff changeset
   444
end"
128
693711a0c702 polished
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   445
"((\"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: 43
diff changeset
   446
108
8bea3f74889d added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
Christian Urban <urbanc@in.tum.de>
parents: 105
diff changeset
   447
  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: 43
diff changeset
   448
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   449
@{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: 43
diff changeset
   450
"let 
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   451
  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: 43
diff changeset
   452
in 
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents: 43
diff changeset
   453
  (OuterParse.$$$ \"|\" -- OuterParse.$$$ \"in\") input
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents: 43
diff changeset
   454
end"
183
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 181
diff changeset
   455
"((\"|\", \"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: 43
diff changeset
   456
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   457
  The parser @{ML "OuterParse.enum s p" for s p} parses a possibly empty 
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
   458
  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: 101
diff changeset
   459
  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: 43
diff changeset
   460
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   461
@{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: 43
diff changeset
   462
"let 
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   463
  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: 43
diff changeset
   464
in 
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents: 43
diff changeset
   465
  (OuterParse.enum \"|\" (OuterParse.$$$ \"in\")) input
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents: 43
diff changeset
   466
end" 
183
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 181
diff changeset
   467
"([\"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: 43
diff changeset
   468
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   469
  @{ML "OuterParse.enum1"} works similarly, except that the parsed list must
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   470
  be non-empty. Note that we had to add a string @{text [quotes] "foo"} at the
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   471
  end of the parsed string, otherwise the parser would have consumed all
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   472
  tokens and then failed with the exception @{text "MORE"}. Like in the
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   473
  previous section, we can avoid this exception using the wrapper @{ML
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   474
  Scan.finite}. This time, however, we have to use the ``stopper-token'' @{ML
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   475
  OuterLex.stopper}. We can write:
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   476
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   477
@{ML_response [display,gray]
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   478
"let 
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   479
  val input = filtered_input \"in | in | in\"
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   480
in 
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   481
  Scan.finite OuterLex.stopper 
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   482
         (OuterParse.enum \"|\" (OuterParse.$$$ \"in\")) input
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   483
end" 
183
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 181
diff changeset
   484
"([\"in\", \"in\", \"in\"], [])"}
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   485
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   486
  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: 52
diff changeset
   487
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   488
*}
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   489
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   490
ML{*fun parse p input = Scan.finite OuterLex.stopper (Scan.error p) input *}
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   491
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   492
text {*
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   493
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   494
  The function @{ML "OuterParse.!!!"} can be used to force termination of the
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   495
  parser in case of a dead end, just like @{ML "Scan.!!"} (see previous section), 
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   496
  except that the error message is fixed to be @{text [quotes] "Outer syntax error"}
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   497
  with a relatively precise description of the failure. For example:
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   498
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   499
@{ML_response_fake [display,gray]
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   500
"let 
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   501
  val input = filtered_input \"in |\"
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   502
  val parse_bar_then_in = OuterParse.$$$ \"|\" -- OuterParse.$$$ \"in\"
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   503
in 
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   504
  parse (OuterParse.!!! parse_bar_then_in) input 
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   505
end"
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   506
"Exception ERROR \"Outer syntax error: keyword \"|\" expected, 
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   507
but keyword in was found\" raised"
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   508
}
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   509
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   510
  \begin{exercise} (FIXME)
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   511
  A type-identifier, for example @{typ "'a"}, is a token of 
54
1783211b3494 tuned; added document antiquotation ML_response_fake_both
Christian Urban <urbanc@in.tum.de>
parents: 53
diff changeset
   512
  kind @{ML "Keyword" in OuterLex}. It can be parsed using 
1783211b3494 tuned; added document antiquotation ML_response_fake_both
Christian Urban <urbanc@in.tum.de>
parents: 53
diff changeset
   513
  the function @{ML OuterParse.type_ident}.
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   514
  \end{exercise}
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   515
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   516
  (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: 52
diff changeset
   517
125
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   518
  Whenever there is a possibility that the processing of user input can fail, 
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   519
  it is a good idea to give as much information about where the error 
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   520
  occured. For this Isabelle can attach positional information to tokens
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   521
  and then thread this information up the processing chain. To see this,
128
693711a0c702 polished
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   522
  modify the function @{ML filtered_input} described earlier to 
41
b11653b11bd3 further progress on the parsing section and tuning on the antiqu's
Christian Urban <urbanc@in.tum.de>
parents: 40
diff changeset
   523
*}
b11653b11bd3 further progress on the parsing section and tuning on the antiqu's
Christian Urban <urbanc@in.tum.de>
parents: 40
diff changeset
   524
125
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   525
ML{*fun filtered_input' str = 
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   526
       filter OuterLex.is_proper (OuterSyntax.scan (Position.line 7) str) *}
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   527
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   528
text {*
125
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   529
  where we pretend the parsed string starts on line 7. An example is
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   530
125
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   531
@{ML_response_fake [display,gray]
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   532
"filtered_input' \"foo \\n bar\""
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   533
"[Token ((\"foo\", ({line=7, end_line=7}, {line=7})), (Ident, \"foo\"), \<dots>),
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   534
 Token ((\"bar\", ({line=8, end_line=8}, {line=8})), (Ident, \"bar\"), \<dots>)]"}
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   535
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   536
  in which the @{text [quotes] "\\n"} causes the second token to be in 
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   537
  line 8.
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   538
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 125
diff changeset
   539
  By using the parser @{ML OuterParse.position} you can decode the positional
125
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   540
  information and return it as part of the parsed input. For example
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   541
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   542
@{ML_response_fake [display,gray]
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   543
"let
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   544
  val input = (filtered_input' \"where\")
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   545
in 
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   546
  parse (OuterParse.position (OuterParse.$$$ \"where\")) input
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   547
end"
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   548
"((\"where\", {line=7, end_line=7}), [])"}
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   549
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   550
  \begin{readmore}
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   551
  The functions related to positions are implemented in the file
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   552
  @{ML_file "Pure/General/position.ML"}.
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   553
  \end{readmore}
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   554
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   555
*}
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   556
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: 43
diff changeset
   557
section {* Parsing Inner Syntax *}
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   558
125
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   559
text {*
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   560
  There is usually no need to write your own parser for parsing inner syntax, that is 
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   561
  for terms and  types: you can just call the pre-defined parsers. Terms can 
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   562
  be parsed using the function @{ML OuterParse.term}. For example:
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   563
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   564
@{ML_response [display,gray]
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   565
"let 
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   566
  val input = OuterSyntax.scan Position.none \"foo\"
44
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents: 43
diff changeset
   567
in 
125
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   568
  OuterParse.term input
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   569
end"
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   570
"(\"\\^E\\^Ftoken\\^Efoo\\^E\\^F\\^E\", [])"}
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   571
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   572
  The function @{ML OuterParse.prop} is similar, except that it gives a different
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   573
  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: 126
diff changeset
   574
  the parsed string, but also some encoded information. You can decode the
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   575
  information with the function @{ML YXML.parse}. For example
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   576
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   577
  @{ML_response [display,gray]
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   578
  "YXML.parse \"\\^E\\^Ftoken\\^Efoo\\^E\\^F\\^E\""
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   579
  "XML.Elem (\"token\", [], [XML.Text \"foo\"])"}
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   580
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
   581
  The result of the decoding is an XML-tree. You can see better what is going on if
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
   582
  you replace @{ML Position.none} by @{ML "Position.line 42"}, say:
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   583
125
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   584
@{ML_response [display,gray]
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   585
"let 
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   586
  val input = OuterSyntax.scan (Position.line 42) \"foo\"
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   587
in 
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   588
  YXML.parse (fst (OuterParse.term input))
125
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   589
end"
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   590
"XML.Elem (\"token\", [(\"line\", \"42\"), (\"end_line\", \"42\")], [XML.Text \"foo\"])"}
125
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   591
  
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
   592
  The positional information is stored as part of an XML-tree so that code 
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
   593
  called later on will be able to give more precise error messages. 
125
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   594
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   595
  \begin{readmore}
128
693711a0c702 polished
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   596
  The functions to do with input and output of XML and YXML are defined 
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   597
  in @{ML_file "Pure/General/xml.ML"} and @{ML_file "Pure/General/yxml.ML"}.
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   598
  \end{readmore}
160
cc9359bfacf4 redefined the functions warning and tracing in order to properly match more antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
   599
  
125
748d9c1a32fb polished parser section
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   600
*}
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   601
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   602
section {* Parsing Specifications\label{sec:parsingspecs} *}
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   603
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   604
text {*
121
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   605
  There are a number of special purpose parsers that help with parsing
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 149
diff changeset
   606
  specifications of function definitions, inductive predicates and so on. In
121
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   607
  Capter~\ref{chp:package}, for example, we will need to parse specifications
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   608
  for inductive predicates of the form:
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   609
*}
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   610
121
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   611
simple_inductive
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   612
  even and odd
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   613
where
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   614
  even0: "even 0"
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   615
| evenS: "odd n \<Longrightarrow> even (Suc n)"
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   616
| oddS: "even n \<Longrightarrow> odd (Suc n)"
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   617
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   618
text {*
121
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   619
  For this we are going to use the parser:
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   620
*}
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   621
121
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   622
ML %linenosgray{*val spec_parser = 
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 125
diff changeset
   623
     OuterParse.fixes -- 
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 125
diff changeset
   624
     Scan.optional 
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 125
diff changeset
   625
       (OuterParse.$$$ "where" |--
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 125
diff changeset
   626
          OuterParse.!!! 
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 125
diff changeset
   627
            (OuterParse.enum1 "|" 
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 125
diff changeset
   628
               (SpecParse.opt_thm_name ":" -- OuterParse.prop))) []*}
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 116
diff changeset
   629
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   630
text {*
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 125
diff changeset
   631
  Note that the parser does not parse the keyword \simpleinductive, even if it is
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 125
diff changeset
   632
  meant to process definitions as shown above. The parser of the keyword 
128
693711a0c702 polished
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   633
  will be given by the infrastructure that will eventually call @{ML spec_parser}.
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 125
diff changeset
   634
  
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 125
diff changeset
   635
124
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 122
diff changeset
   636
  To see what the parser returns, let us parse the string corresponding to the 
121
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   637
  definition of @{term even} and @{term odd}:
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   638
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   639
@{ML_response [display,gray]
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   640
"let
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   641
  val input = filtered_input
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   642
     (\"even and odd \" ^  
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   643
      \"where \" ^
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   644
      \"  even0[intro]: \\\"even 0\\\" \" ^ 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   645
      \"| evenS[intro]: \\\"odd n \<Longrightarrow> even (Suc n)\\\" \" ^ 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   646
      \"| oddS[intro]:  \\\"even n \<Longrightarrow> odd (Suc n)\\\"\")
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   647
in
120
c39f83d8daeb some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents: 116
diff changeset
   648
  parse spec_parser input
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   649
end"
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   650
"(([(even, NONE, NoSyn), (odd, NONE, NoSyn)],
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   651
     [((even0,\<dots>), \"\\^E\\^Ftoken\\^Eeven 0\\^E\\^F\\^E\"),
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   652
      ((evenS,\<dots>), \"\\^E\\^Ftoken\\^Eodd n \<Longrightarrow> even (Suc n)\\^E\\^F\\^E\"),
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   653
      ((oddS,\<dots>), \"\\^E\\^Ftoken\\^Eeven n \<Longrightarrow> odd (Suc n)\\^E\\^F\\^E\")]), [])"}
121
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   654
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   655
  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: 183
diff changeset
   656
  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: 183
diff changeset
   657
  rules where every rule has optionally a name and an attribute.
121
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   658
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   659
  The function @{ML OuterParse.fixes} in Line 2 of the parser reads an 
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   660
  \isacommand{and}-separated 
124
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 122
diff changeset
   661
  list of variables that can include optional type annotations and syntax translations. 
121
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   662
  For example:\footnote{Note that in the code we need to write 
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   663
  @{text "\\\"int \<Rightarrow> bool\\\""} in order to properly escape the double quotes
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   664
  in the compound type.}
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   665
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   666
@{ML_response [display,gray]
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   667
"let
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   668
  val input = filtered_input 
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   669
        \"foo::\\\"int \<Rightarrow> bool\\\" and bar::nat (\\\"BAR\\\" 100) and blonk\"
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   670
in
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   671
   parse OuterParse.fixes input
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   672
end"
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   673
"([(foo, SOME \"\\^E\\^Ftoken\\^Eint \<Rightarrow> bool\\^E\\^F\\^E\", NoSyn), 
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   674
  (bar, SOME \"\\^E\\^Ftoken\\^Enat\\^E\\^F\\^E\", Mixfix (\"BAR\", [], 100)), 
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   675
  (blonk, NONE, NoSyn)],[])"}  
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   676
*}
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   677
121
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   678
text {*
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 149
diff changeset
   679
  Whenever types are given, they are stored in the @{ML SOME}s. The types are
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 149
diff changeset
   680
  not yet used to type the variables: this must be done by type-inference later
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
   681
  on. Since types are part of the inner syntax they are strings with some
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
   682
  encoded information (see previous section). If a syntax translation is
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
   683
  present for a variable, then it is stored in the @{ML Mixfix} datastructure;
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
   684
  no syntax translation is indicated by @{ML NoSyn}.
121
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   685
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   686
  \begin{readmore}
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   687
  The datastructre for sytax annotations is defined in @{ML_file "Pure/Syntax/mixfix.ML"}.
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   688
  \end{readmore}
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   689
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   690
  Lines 3 to 7 in the function @{ML spec_parser} implement the parser for a
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
   691
  list of introduction rules, that is propositions with theorem
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
   692
  annotations. The introduction rules are propositions parsed by @{ML
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
   693
  OuterParse.prop}. However, they can include an optional theorem name plus
121
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   694
  some attributes. For example
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   695
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   696
@{ML_response [display,gray] "let 
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   697
  val input = filtered_input \"foo_lemma[intro,dest!]:\"
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   698
  val ((name, attrib), _) = parse (SpecParse.thm_name \":\") input 
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   699
in 
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   700
  (name, map Args.dest_src attrib)
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   701
end" "(foo_lemma, [((\"intro\", []), \<dots>), ((\"dest\", [\<dots>]), \<dots>)])"}
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   702
 
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   703
  The function @{ML opt_thm_name in SpecParse} is the ``optional'' variant of
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
   704
  @{ML thm_name in SpecParse}. Theorem names can contain attributes. The name 
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
   705
  has to end with @{text [quotes] ":"}---see the argument of 
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   706
  the function @{ML SpecParse.opt_thm_name} in Line 7.
121
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   707
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   708
  \begin{readmore}
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   709
  Attributes and arguments are implemented in the files @{ML_file "Pure/Isar/attrib.ML"} 
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   710
  and @{ML_file "Pure/Isar/args.ML"}.
26e5b41faa74 polishing
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   711
  \end{readmore}
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   712
*}
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   713
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   714
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: 60
diff changeset
   715
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   716
text {*
188
8939b8fd8603 added lots of FIXMES
Christian Urban <urbanc@in.tum.de>
parents: 186
diff changeset
   717
  (FIXME: update to the right command setup)
8939b8fd8603 added lots of FIXMES
Christian Urban <urbanc@in.tum.de>
parents: 186
diff changeset
   718
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   719
  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: 67
diff changeset
   720
  need to be implemented. While this is not difficult on the ML-level,
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   721
  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: 60
diff changeset
   722
  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: 60
diff changeset
   723
  will explain in this section.
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   724
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   725
  To keep things simple, let us start with a ``silly'' command that does nothing 
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   726
  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: 101
diff changeset
   727
  defined as:
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   728
*}
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   729
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   730
ML{*let
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   731
  val do_nothing = Scan.succeed (Toplevel.theory I)
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   732
  val kind = OuterKeyword.thy_decl
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   733
in
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   734
  OuterSyntax.command "foobar" "description of foobar" kind do_nothing
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   735
end *}
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   736
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   737
text {*
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   738
  The crucial function @{ML OuterSyntax.command} expects a name for the command, a
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   739
  short description, a kind indicator (which we will explain later on more thoroughly) and a
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   740
  parser producing a top-level transition function (its purpose will also explained
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   741
  later). 
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   742
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   743
  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: 67
diff changeset
   744
  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: 67
diff changeset
   745
  recognise \isacommand{foobar} as a command. Such a keyword file can be
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   746
  generated with the command-line:
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   747
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   748
  @{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: 60
diff changeset
   749
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   750
  The option @{text "-k foobar"} indicates which postfix the name of the keyword file 
80
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   751
  will be assigned. In the case above the file will be named @{text
86
fdb9ea86c2a3 polished
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
   752
  "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: 67
diff changeset
   753
  present (in order to extract the keywords from them). To generate these log
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   754
  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: 67
diff changeset
   755
  @{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: 67
diff changeset
   756
  complete code.
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   757
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   758
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   759
  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   760
  \begin{figure}[t]
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   761
  \begin{graybox}\small
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   762
  \isacommand{theory}~@{text Command}\\
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   763
  \isacommand{imports}~@{text Main}\\
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   764
  \isacommand{begin}\\
85
b02904872d6b better handling of {* and *}
Christian Urban <urbanc@in.tum.de>
parents: 81
diff changeset
   765
  \isacommand{ML}~@{text "\<verbopen>"}\\
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   766
  @{ML
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   767
"let
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   768
  val do_nothing = Scan.succeed (Toplevel.theory I)
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   769
  val kind = OuterKeyword.thy_decl
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   770
in
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   771
  OuterSyntax.command \"foobar\" \"description of foobar\" kind do_nothing
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   772
end"}\\
85
b02904872d6b better handling of {* and *}
Christian Urban <urbanc@in.tum.de>
parents: 81
diff changeset
   773
  @{text "\<verbclose>"}\\
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   774
  \isacommand{end}
80
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   775
  \end{graybox}
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   776
  \caption{\small The file @{text "Command.thy"} is necessary for generating a log 
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   777
  file. This log file enables Isabelle to generate a keyword file containing 
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   778
  the command \isacommand{foobar}.\label{fig:commandtheory}}
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   779
  \end{figure}
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   780
  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   781
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   782
  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: 67
diff changeset
   783
  @{text "Pure"}, @{text "HOL"} and @{text "Pure-ProofGeneral"}, as well as
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   784
  the log file for the theory @{text "Command.thy"}, which contains the new
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   785
  \isacommand{foobar}-command. If you target other logics besides HOL, such
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   786
  as Nominal or ZF, then you need to adapt the log files appropriately.
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   787
  
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   788
  @{text Pure} and @{text HOL} are usually compiled during the installation of
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   789
  Isabelle. So log files for them should be already available. If not, then
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   790
  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: 101
diff changeset
   791
  distribution.
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   792
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   793
  @{text [display] 
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   794
"$ ./build -m \"Pure\"
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   795
$ ./build -m \"HOL\""}
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   796
  
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   797
  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: 60
diff changeset
   798
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   799
  @{text [display] "$ ./build -m \"Pure-ProofGeneral\" \"Pure\""}
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   800
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   801
  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: 101
diff changeset
   802
  with:
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   803
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   804
  @{text [display] "$ isabelle mkdir FoobarCommand"}
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   805
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   806
  This generates a directory containing the files: 
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   807
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   808
  @{text [display] 
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   809
"./IsaMakefile
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   810
./FoobarCommand/ROOT.ML
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   811
./FoobarCommand/document
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   812
./FoobarCommand/document/root.tex"}
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   813
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   814
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   815
  You need to copy the file @{text "Command.thy"} into the directory @{text "FoobarCommand"}
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   816
  and add the line 
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   817
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   818
  @{text [display] "use_thy \"Command\";"} 
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   819
  
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   820
  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: 60
diff changeset
   821
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   822
  @{text [display] "$ isabelle make"}
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   823
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   824
  If the compilation succeeds, you have finally created all the necessary log files. 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   825
  They are stored in the directory 
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   826
  
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   827
  @{text [display]  "~/.isabelle/heaps/Isabelle2008/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: 60
diff changeset
   828
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   829
  or something similar depending on your Isabelle distribution and architecture.
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   830
  One quick way to assign a shell variable to this directory is by typing
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   831
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   832
  @{text [display] "$ ISABELLE_LOGS=\"$(isabelle getenv -b ISABELLE_OUTPUT)\"/log"}
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   833
 
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 149
diff changeset
   834
  on the Unix prompt. If you now type @{text "ls $ISABELLE_LOGS"}, then the 
128
693711a0c702 polished
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   835
  directory should include the files:
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   836
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   837
  @{text [display] 
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   838
"Pure.gz
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   839
HOL.gz
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   840
Pure-ProofGeneral.gz
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   841
HOL-FoobarCommand.gz"} 
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   842
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   843
  From them you can create the keyword files. Assuming the name 
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   844
  of the directory is in  @{text "$ISABELLE_LOGS"},
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   845
  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: 60
diff changeset
   846
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   847
@{text [display]
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   848
"$ isabelle keywords -k foobar 
80
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   849
   $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: 60
diff changeset
   850
80
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   851
  The result is the file @{text "isar-keywords-foobar.el"}. It should contain
86
fdb9ea86c2a3 polished
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
   852
  the string @{text "foobar"} twice.\footnote{To see whether things are fine, check
80
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   853
  that @{text "grep foobar"} on this file returns something
86
fdb9ea86c2a3 polished
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
   854
  non-empty.}  This keyword file needs to
80
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   855
  be copied into the directory @{text "~/.isabelle/etc"}. To make Isabelle
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   856
  aware of this keyword file, you have to start Isabelle with the option @{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: 101
diff changeset
   857
  "-k foobar"}, that is:
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   858
80
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   859
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   860
  @{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: 60
diff changeset
   861
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   862
  If you now build a theory on top of @{text "Command.thy"}, 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   863
  then the command \isacommand{foobar} can be used. 
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   864
  Similarly with any other new command. 
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   865
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   866
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   867
  At the moment \isacommand{foobar} is not very useful. Let us refine it a bit 
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   868
  next by letting it take a proposition as argument and printing this proposition 
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   869
  inside the tracing buffer. 
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   870
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   871
  The crucial part of a command is the function that determines the behaviour
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   872
  of the command. In the code above we used a ``do-nothing''-function, which
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   873
  because of @{ML Scan.succeed} does not parse any argument, but immediately
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   874
  returns the simple toplevel function @{ML "Toplevel.theory I"}. We can
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   875
  replace this code by a function that first parses a proposition (using the
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   876
  parser @{ML OuterParse.prop}), then prints out the tracing
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   877
  information (using a new top-level function @{text trace_top_lvl}) 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: 101
diff changeset
   878
  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: 67
diff changeset
   879
*}
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   880
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   881
ML{*let
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   882
  fun trace_top_lvl str = 
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   883
     Toplevel.theory (fn thy => (tracing str; thy))
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   884
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   885
  val trace_prop = OuterParse.prop >> trace_top_lvl
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   886
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   887
  val kind = OuterKeyword.thy_decl
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   888
in
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   889
  OuterSyntax.command "foobar" "traces a proposition" kind trace_prop
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   890
end *}
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   891
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   892
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: 101
diff changeset
   893
  Now you can type
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   894
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   895
  \begin{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   896
  \isacommand{foobar}~@{text [quotes] "True \<and> False"}\\
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   897
  @{text "> \"True \<and> False\""}
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   898
  \end{isabelle}
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   899
  
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   900
  and see the proposition in the tracing buffer.  
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   901
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   902
  Note that so far we used @{ML thy_decl in OuterKeyword} as the kind indicator
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   903
  for the command.  This means that the command finishes as soon as the
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   904
  arguments are processed. Examples of this kind of commands are
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   905
  \isacommand{definition} and \isacommand{declare}.  In other cases,
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   906
  commands are expected to parse some arguments, for example a proposition,
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   907
  and then ``open up'' a proof in order to prove the proposition (for example
86
fdb9ea86c2a3 polished
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
   908
  \isacommand{lemma}) or prove some other properties (for example
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   909
  \isacommand{function}). To achieve this kind of behaviour, you have to use the kind
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
   910
  indicator @{ML thy_goal in OuterKeyword}.  Note, however, once you change the 
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
   911
  ``kind'' of a command from @{ML thy_decl in OuterKeyword} to @{ML thy_goal in OuterKeyword} 
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   912
  then the keyword file needs to be re-created!
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   913
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   914
  Below we change \isacommand{foobar} so that it takes a proposition 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: 101
diff changeset
   915
  argument and then starts a proof in order to prove it. Therefore in Line 13, 
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   916
  we set the kind indicator to @{ML thy_goal in OuterKeyword}.
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   917
*}
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   918
114
13fd0a83d3c3 properly handled linenumbers in ML-text and Isar-proofs
Christian Urban <urbanc@in.tum.de>
parents: 108
diff changeset
   919
ML%linenosgray{*let
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   920
  fun set_up_thm str ctxt =
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   921
    let
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   922
      val prop = Syntax.read_prop ctxt str
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   923
    in
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   924
      Proof.theorem_i NONE (K I) [[(prop,[])]] ctxt
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   925
    end;
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   926
  
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   927
  val prove_prop = OuterParse.prop >>  
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   928
      (fn str => Toplevel.print o 
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   929
                    Toplevel.local_theory_to_proof NONE (set_up_thm str))
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   930
  
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   931
  val kind = OuterKeyword.thy_goal
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   932
in
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   933
  OuterSyntax.command "foobar" "proving a proposition" kind prove_prop
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   934
end *}
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   935
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   936
text {*
86
fdb9ea86c2a3 polished
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
   937
  The function @{text set_up_thm} 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: 101
diff changeset
   938
  proved) and a context as argument.  The context is necessary in order to be able to use
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   939
  @{ML Syntax.read_prop}, which converts a string into a proper proposition.
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   940
  In Line 6 the function @{ML Proof.theorem_i} starts the proof for the
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   941
  proposition. Its argument @{ML NONE} stands for a locale (which we chose to
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   942
  omit); the argument @{ML "(K I)"} stands for a function that determines what
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   943
  should be done with the theorem once it is proved (we chose to just forget
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   944
  about it). Lines 9 to 11 contain the parser for the proposition.
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   945
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   946
  If you now type \isacommand{foobar}~@{text [quotes] "True \<and> True"}, you obtain the following 
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   947
  proof state:
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   948
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   949
  \begin{isabelle}
80
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   950
  \isacommand{foobar}~@{text [quotes] "True \<and> True"}\\
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   951
  @{text "goal (1 subgoal):"}\\
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   952
  @{text "1. True \<and> True"}
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   953
  \end{isabelle}
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   954
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
   955
  and you can build the proof
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   956
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   957
  \begin{isabelle}
80
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   958
  \isacommand{foobar}~@{text [quotes] "True \<and> True"}\\
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   959
  \isacommand{apply}@{text "(rule conjI)"}\\
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   960
  \isacommand{apply}@{text "(rule TrueI)+"}\\
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   961
  \isacommand{done}
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   962
  \end{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   963
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 133
diff changeset
   964
 
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   965
  
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   966
  (FIXME What do @{ML "Toplevel.theory"} 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   967
  @{ML "Toplevel.print"} 
128
693711a0c702 polished
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   968
  @{ML Toplevel.local_theory} do?)
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   969
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   970
  (FIXME read a name and show how to store theorems)
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   971
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   972
*}
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   973
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   974
section {* Methods *}
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   975
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   976
text {*
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   977
  Methods are a central concept in Isabelle. They are the ones you use for example
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   978
  in \isacommand{apply}. To print out all currently known methods you can use the 
192
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   979
  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: 177
diff changeset
   980
*}
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   981
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   982
print_methods
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   983
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   984
text {*
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   985
  An example of a very simple method is the following code.
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   986
*}
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   987
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   988
method_setup %gray foobar_meth = 
181
5baaabe1ab92 updated to new method_setup
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   989
 {* Scan.succeed
5baaabe1ab92 updated to new method_setup
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   990
      (K (SIMPLE_METHOD ((etac @{thm conjE} THEN' rtac @{thm conjI}) 1))) *}
5baaabe1ab92 updated to new method_setup
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   991
         "foobar method"
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   992
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   993
text {*
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   994
  It defines the method @{text foobar_meth}, which takes no arguments (therefore the
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   995
  parser @{ML Scan.succeed}) and 
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   996
  only applies the tactic @{thm [source] conjE} and then @{thm [source] conjI}.
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   997
  This method can be used in the following proof
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   998
*}
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   999
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
  1000
lemma shows "A \<and> B \<Longrightarrow> C \<and> D"
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
  1001
apply(foobar_meth)
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
  1002
txt {*
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
  1003
  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: 177
diff changeset
  1004
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
  1005
  \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: 177
diff changeset
  1006
  @{subgoals}
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
  1007
  \end{minipage} *}
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
  1008
(*<*)oops(*>*)
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
  1009
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
  1010
text {*
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
  1011
  (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: 183
diff changeset
  1012
*}
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
  1013
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
  1014
(*<*)
80
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
  1015
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
  1016
chapter {* Parsing *}
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
  1017
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
  1018
text {*
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
  1019
4
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1020
  Lots of Standard ML code is given in this document, for various reasons,
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1021
  including:
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1022
  \begin{itemize}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1023
  \item direct quotation of code found in the Isabelle source files,
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1024
  or simplified versions of such code
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1025
  \item identifiers found in the Isabelle source code, with their types 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1026
  (or specialisations of their types)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1027
  \item code examples, which can be run by the reader, to help illustrate the
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1028
  behaviour of functions found in the Isabelle source code
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1029
  \item ancillary functions, not from the Isabelle source code, 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1030
  which enable the reader to run relevant code examples
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1031
  \item type abbreviations, which help explain the uses of certain functions
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1032
  \end{itemize}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1033
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1034
*}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1035
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1036
section {* Parsing Isar input *}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1037
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1038
text {*
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1039
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1040
  The typical parsing function has the type
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1041
  \texttt{'src -> 'res * 'src}, with input  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1042
  of type \texttt{'src}, returning a result 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1043
  of type \texttt{'res}, which is (or is derived from) the first part of the
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1044
  input, and also returning the remainder of the input.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1045
  (In the common case, when it is clear what the ``remainder of the input''
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1046
  means, we will just say that the functions ``returns'' the
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1047
  value of type \texttt{'res}). 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1048
  An exception is raised if an appropriate value 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1049
  cannot be produced from the input.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1050
  A range of exceptions can be used to identify different reasons 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1051
  for the failure of a parse.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1052
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1053
  This contrasts the standard parsing function in Standard ML,
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1054
  which is of type 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1055
  \texttt{type ('res, 'src) reader = 'src -> ('res * 'src) option};
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1056
  (for example, \texttt{List.getItem} and \texttt{Substring.getc}).
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1057
  However, much of the discussion at 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1058
  FIX file:/home/jeremy/html/ml/SMLBasis/string-cvt.html
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1059
  is relevant.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1060
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1061
  Naturally one may convert between the two different sorts of parsing functions
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1062
  as follows:
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1063
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1064
  open StringCvt ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1065
  type ('res, 'src) ex_reader = 'src -> 'res * 'src
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
  1066
  ex_reader : ('res, 'src) reader -> ('res, 'src) ex_reader 
4
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1067
  fun ex_reader rdr src = Option.valOf (rdr src) ;
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
  1068
  reader : ('res, 'src) ex_reader -> ('res, 'src) reader 
4
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1069
  fun reader exrdr src = SOME (exrdr src) handle _ => NONE ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1070
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1071
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1072
*}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1073
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1074
section{* The \texttt{Scan} structure *}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1075
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1076
text {* 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1077
  The source file is \texttt{src/General/scan.ML}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1078
  This structure provides functions for using and combining parsing functions
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1079
  of the type \texttt{'src -> 'res * 'src}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1080
  Three exceptions are used:
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1081
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1082
  exception MORE of string option;  (*need more input (prompt)*)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1083
  exception FAIL of string option;  (*try alternatives (reason of failure)*)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1084
  exception ABORT of string;        (*dead end*)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1085
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1086
  Many functions in this structure (generally those with names composed of
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1087
  symbols) are declared as infix.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1088
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1089
  Some functions from that structure are
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1090
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1091
  |-- : ('src -> 'res1 * 'src') * ('src' -> 'res2 * 'src'') ->
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1092
  'src -> 'res2 * 'src''
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1093
  --| : ('src -> 'res1 * 'src') * ('src' -> 'res2 * 'src'') ->
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1094
  'src -> 'res1 * 'src''
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1095
  -- : ('src -> 'res1 * 'src') * ('src' -> 'res2 * 'src'') ->
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1096
  'src -> ('res1 * 'res2) * 'src''
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1097
  ^^ : ('src -> string * 'src') * ('src' -> string * 'src'') ->
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1098
  'src -> string * 'src''
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1099
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1100
  These functions parse a result off the input source twice.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1101
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1102
  \texttt{|--} and \texttt{--|} 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1103
  return the first result and the second result, respectively.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1104
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1105
  \texttt{--} returns both.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1106
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1107
  \verb|^^| returns the result of concatenating the two results
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1108
  (which must be strings).
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1109
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1110
  Note how, although the types 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1111
  \texttt{'src}, \texttt{'src'} and \texttt{'src''} will normally be the same,
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1112
  the types as shown help suggest the behaviour of the functions.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1113
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1114
  :-- : ('src -> 'res1 * 'src') * ('res1 -> 'src' -> 'res2 * 'src'') ->
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1115
  'src -> ('res1 * 'res2) * 'src''
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1116
  :|-- : ('src -> 'res1 * 'src') * ('res1 -> 'src' -> 'res2 * 'src'') ->
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1117
  'src -> 'res2 * 'src''
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1118
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1119
  These are similar to \texttt{|--} and \texttt{--|},
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1120
  except that the second parsing function can depend on the result of the first.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1121
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1122
  >> : ('src -> 'res1 * 'src') * ('res1 -> 'res2) -> 'src -> 'res2 * 'src'
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1123
  || : ('src -> 'res_src) * ('src -> 'res_src) -> 'src -> 'res_src
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1124
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1125
  \texttt{p >> f} applies a function \texttt{f} to the result of a parse.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1126
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1127
  \texttt{||} tries a second parsing function if the first one
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1128
  fails by raising an exception of the form \texttt{FAIL \_}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1129
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1130
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1131
  succeed : 'res -> ('src -> 'res * 'src) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1132
  fail : ('src -> 'res_src) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1133
  !! : ('src * string option -> string) -> 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1134
  ('src -> 'res_src) -> ('src -> 'res_src) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1135
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1136
  \texttt{succeed r} returns \texttt{r}, with the input unchanged.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1137
  \texttt{fail} always fails, raising exception \texttt{FAIL NONE}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1138
  \texttt{!! f} only affects the failure mode, turning a failure that
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1139
  raises \texttt{FAIL \_} into a failure that raises \texttt{ABORT ...}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1140
  This is used to prevent recovery from the failure ---
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1141
  thus, in \texttt{!! parse1 || parse2}, if \texttt{parse1} fails, 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1142
  it won't recover by trying \texttt{parse2}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1143
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1144
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1145
  one : ('si -> bool) -> ('si list -> 'si * 'si list) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1146
  some : ('si -> 'res option) -> ('si list -> 'res * 'si list) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1147
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1148
  These require the input to be a list of items:
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1149
  they fail, raising \texttt{MORE NONE} if the list is empty.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1150
  On other failures they raise \texttt{FAIL NONE} 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1151
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1152
  \texttt{one p} takes the first
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1153
  item from the list if it satisfies \texttt{p}, otherwise fails.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1154
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1155
  \texttt{some f} takes the first
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1156
  item from the list and applies \texttt{f} to it, failing if this returns
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1157
  \texttt{NONE}.  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1158
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1159
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1160
  many : ('si -> bool) -> 'si list -> 'si list * 'si list ; 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1161
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1162
  \texttt{many p} takes items from the input until it encounters one 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1163
  which does not satisfy \texttt{p}.  If it reaches the end of the input
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1164
  it fails, raising \texttt{MORE NONE}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1165
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1166
  \texttt{many1} (with the same type) fails if the first item 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1167
  does not satisfy \texttt{p}.  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1168
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1169
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1170
  option : ('src -> 'res * 'src) -> ('src -> 'res option * 'src)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1171
  optional : ('src -> 'res * 'src) -> 'res -> ('src -> 'res * 'src)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1172
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1173
  \texttt{option}: 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1174
  where the parser \texttt{f} succeeds with result \texttt{r} 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1175
  or raises \texttt{FAIL \_},
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1176
  \texttt{option f} gives the result \texttt{SOME r} or \texttt{NONE}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1177
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1178
  \texttt{optional}: if parser \texttt{f} fails by raising \texttt{FAIL \_},
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1179
  \texttt{optional f default} provides the result \texttt{default}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1180
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1181
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1182
  repeat : ('src -> 'res * 'src) -> 'src -> 'res list * 'src
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1183
  repeat1 : ('src -> 'res * 'src) -> 'src -> 'res list * 'src
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1184
  bulk : ('src -> 'res * 'src) -> 'src -> 'res list * 'src 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1185
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1186
  \texttt{repeat f} repeatedly parses an item off the remaining input until 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1187
  \texttt{f} fails with \texttt{FAIL \_}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1188
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1189
  \texttt{repeat1} is as for \texttt{repeat}, but requires at least one
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1190
  successful parse.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1191
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1192
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1193
  lift : ('src -> 'res * 'src) -> ('ex * 'src -> 'res * ('ex * 'src))
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1194
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1195
  \texttt{lift} changes the source type of a parser by putting in an extra
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1196
  component \texttt{'ex}, which is ignored in the parsing.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1197
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1198
  The \texttt{Scan} structure also provides the type \texttt{lexicon}, 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1199
  HOW DO THEY WORK ?? TO BE COMPLETED
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1200
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1201
  dest_lexicon: lexicon -> string list ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1202
  make_lexicon: string list list -> lexicon ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1203
  empty_lexicon: lexicon ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1204
  extend_lexicon: string list list -> lexicon -> lexicon ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1205
  merge_lexicons: lexicon -> lexicon -> lexicon ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1206
  is_literal: lexicon -> string list -> bool ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1207
  literal: lexicon -> string list -> string list * string list ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1208
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1209
  Two lexicons, for the commands and keywords, are stored and can be retrieved
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1210
  by:
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1211
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1212
  val (command_lexicon, keyword_lexicon) = OuterSyntax.get_lexicons () ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1213
  val commands = Scan.dest_lexicon command_lexicon ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1214
  val keywords = Scan.dest_lexicon keyword_lexicon ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1215
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1216
*}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1217
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1218
section{* The \texttt{OuterLex} structure *}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1219
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1220
text {*
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1221
  The source file is @{text "src/Pure/Isar/outer_lex.ML"}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1222
  In some other source files its name is abbreviated:
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1223
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1224
  structure T = OuterLex;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1225
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1226
  This structure defines the type \texttt{token}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1227
  (The types
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1228
  \texttt{OuterLex.token},
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1229
  \texttt{OuterParse.token} and
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1230
  \texttt{SpecParse.token} are all the same).
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1231
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1232
  Input text is split up into tokens, and the input source type for many parsing
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1233
  functions is \texttt{token list}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1234
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1235
  The datatype definition (which is not published in the signature) is
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1236
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1237
  datatype token = Token of Position.T * (token_kind * string);
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1238
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1239
  but here are some runnable examples for viewing tokens: 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1240
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1241
*}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1242
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1243
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1244
4
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1245
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
  1246
ML{*
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1247
  val toks = OuterSyntax.scan Position.none
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1248
   "theory,imports;begin x.y.z apply ?v1 ?'a 'a -- || 44 simp (* xx *) { * fff * }" ;
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1249
*}
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1250
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
  1251
ML{*
4
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1252
  print_depth 20 ;
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1253
*}
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1254
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
  1255
ML{*
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1256
  map OuterLex.text_of toks ;
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1257
*}
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1258
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
  1259
ML{*
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1260
  val proper_toks = filter OuterLex.is_proper toks ;
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1261
*}  
4
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1262
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
  1263
ML{*
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1264
  map OuterLex.kind_of proper_toks 
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1265
*}
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1266
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
  1267
ML{*
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1268
  map OuterLex.unparse proper_toks ;
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1269
*}
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1270
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
  1271
ML{*
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1272
  OuterLex.stopper
4
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1273
*}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1274
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1275
text {*
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1276
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1277
  The function \texttt{is\_proper : token -> bool} identifies tokens which are
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1278
  not white space or comments: many parsing functions assume require spaces or
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1279
  comments to have been filtered out.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1280
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1281
  There is a special end-of-file token:
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1282
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1283
  val (tok_eof : token, is_eof : token -> bool) = T.stopper ; 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1284
  (* end of file token *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1285
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1286
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1287
*}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1288
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1289
section {* The \texttt{OuterParse} structure *}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1290
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1291
text {*
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1292
  The source file is \texttt{src/Pure/Isar/outer\_parse.ML}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1293
  In some other source files its name is abbreviated:
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1294
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1295
  structure P = OuterParse;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1296
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1297
  Here the parsers use \texttt{token list} as the input source type. 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1298
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1299
  Some of the parsers simply select the first token, provided that it is of the
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1300
  right kind (as returned by \texttt{T.kind\_of}): these are 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1301
  \texttt{ command, keyword, short\_ident, long\_ident, sym\_ident, term\_var,
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1302
  type\_ident, type\_var, number, string, alt\_string, verbatim, sync, eof}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1303
  Others select the first token, provided that it is one of several kinds,
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1304
  (eg, \texttt{name, xname, text, typ}).
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1305
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1306
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1307
  type 'a tlp = token list -> 'a * token list ; (* token list parser *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1308
  $$$ : string -> string tlp
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1309
  nat : int tlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1310
  maybe : 'a tlp -> 'a option tlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1311
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1312
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1313
  \texttt{\$\$\$ s} returns the first token,
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1314
  if it equals \texttt{s} \emph{and} \texttt{s} is a keyword.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1315
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1316
  \texttt{nat} returns the first token, if it is a number, and evaluates it.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1317
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1318
  \texttt{maybe}: if \texttt{p} returns \texttt{r}, 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1319
  then \texttt{maybe p} returns \texttt{SOME r} ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1320
  if the first token is an underscore, it returns \texttt{NONE}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1321
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1322
  A few examples:
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1323
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1324
  P.list : 'a tlp -> 'a list tlp ; (* likewise P.list1 *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1325
  P.and_list : 'a tlp -> 'a list tlp ; (* likewise P.and_list1 *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1326
  val toks : token list = OuterSyntax.scan "44 ,_, 66,77" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1327
  val proper_toks = List.filter T.is_proper toks ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1328
  P.list P.nat toks ; (* OK, doesn't recognize white space *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1329
  P.list P.nat proper_toks ; (* fails, doesn't recognize what follows ',' *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1330
  P.list (P.maybe P.nat) proper_toks ; (* fails, end of input *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1331
  P.list (P.maybe P.nat) (proper_toks @ [tok_eof]) ; (* OK *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1332
  val toks : token list = OuterSyntax.scan "44 and 55 and 66 and 77" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1333
  P.and_list P.nat (List.filter T.is_proper toks @ [tok_eof]) ; (* ??? *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1334
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1335
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1336
  The following code helps run examples:
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1337
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1338
  fun parse_str tlp str = 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1339
  let val toks : token list = OuterSyntax.scan str ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1340
  val proper_toks = List.filter T.is_proper toks @ [tok_eof] ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1341
  val (res, rem_toks) = tlp proper_toks ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1342
  val rem_str = String.concat
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1343
  (Library.separate " " (List.map T.unparse rem_toks)) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1344
  in (res, rem_str) end ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1345
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1346
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1347
  Some examples from \texttt{src/Pure/Isar/outer\_parse.ML}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1348
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1349
  val type_args =
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1350
  type_ident >> Library.single ||
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1351
  $$$ "(" |-- !!! (list1 type_ident --| $$$ ")") ||
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1352
  Scan.succeed [];
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1353
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1354
  There are three ways parsing a list of type arguments can succeed.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1355
  The first line reads a single type argument, and turns it into a singleton
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1356
  list.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1357
  The second line reads "(", and then the remainder, ignoring the "(" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1358
  the remainder consists of a list of type identifiers (at least one),
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1359
  and then a ")" which is also ignored.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1360
  The \texttt{!!!} ensures that if the parsing proceeds this far and then fails,
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1361
  it won't try the third line (see the description of \texttt{Scan.!!}).
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1362
  The third line consumes no input and returns the empty list.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1363
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1364
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1365
  fun triple2 (x, (y, z)) = (x, y, z);
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1366
  val arity = xname -- ($$$ "::" |-- !!! (
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1367
  Scan.optional ($$$ "(" |-- !!! (list1 sort --| $$$ ")")) []
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1368
  -- sort)) >> triple2;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1369
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1370
  The parser \texttt{arity} reads a typename $t$, then ``\texttt{::}'' (which is
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1371
  ignored), then optionally a list $ss$ of sorts and then another sort $s$.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1372
  The result $(t, (ss, s))$ is transformed by \texttt{triple2} to $(t, ss, s)$.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1373
  The second line reads the optional list of sorts:
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1374
  it reads first ``\texttt{(}'' and last ``\texttt{)}'', which are both ignored,
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1375
  and between them a comma-separated list of sorts.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1376
  If this list is absent, the default \texttt{[]} provides the list of sorts.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1377
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1378
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1379
  parse_str P.type_args "('a, 'b) ntyp" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1380
  parse_str P.type_args "'a ntyp" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1381
  parse_str P.type_args "ntyp" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1382
  parse_str P.arity "ty :: tycl" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1383
  parse_str P.arity "ty :: (tycl1, tycl2) tycl" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1384
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1385
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1386
*}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1387
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1388
section {* The \texttt{SpecParse} structure *}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1389
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1390
text {*
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1391
  The source file is \texttt{src/Pure/Isar/spec\_parse.ML}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1392
  This structure contains token list parsers for more complicated values.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1393
  For example, 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1394
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1395
  open SpecParse ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1396
  attrib : Attrib.src tok_rdr ; 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1397
  attribs : Attrib.src list tok_rdr ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1398
  opt_attribs : Attrib.src list tok_rdr ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1399
  xthm : (thmref * Attrib.src list) tok_rdr ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1400
  xthms1 : (thmref * Attrib.src list) list tok_rdr ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1401
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1402
  parse_str attrib "simp" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1403
  parse_str opt_attribs "hello" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1404
  val (ass, "") = parse_str attribs "[standard, xxxx, simp, intro, OF sym]" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1405
  map Args.dest_src ass ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1406
  val (asrc, "") = parse_str attrib "THEN trans [THEN sym]" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1407
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1408
  parse_str xthm "mythm [attr]" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1409
  parse_str xthms1 "thm1 [attr] thms2" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1410
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1411
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1412
  As you can see, attributes are described using types of the \texttt{Args}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1413
  structure, described below.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1414
*}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1415
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1416
section{* The \texttt{Args} structure *}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1417
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1418
text {*
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1419
  The source file is \texttt{src/Pure/Isar/args.ML}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1420
  The primary type of this structure is the \texttt{src} datatype;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1421
  the single constructors not published in the signature, but 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1422
  \texttt{Args.src} and \texttt{Args.dest\_src}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1423
  are in fact the constructor and destructor functions.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1424
  Note that the types \texttt{Attrib.src} and \texttt{Method.src}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1425
  are in fact \texttt{Args.src}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1426
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1427
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1428
  src : (string * Args.T list) * Position.T -> Args.src ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1429
  dest_src : Args.src -> (string * Args.T list) * Position.T ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1430
  Args.pretty_src : Proof.context -> Args.src -> Pretty.T ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1431
  fun pr_src ctxt src = Pretty.string_of (Args.pretty_src ctxt src) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1432
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1433
  val thy = ML_Context.the_context () ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1434
  val ctxt = ProofContext.init thy ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1435
  map (pr_src ctxt) ass ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1436
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1437
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1438
  So an \texttt{Args.src} consists of the first word, then a list of further 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1439
  ``arguments'', of type \texttt{Args.T}, with information about position in the
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1440
  input.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1441
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1442
  (* how an Args.src is parsed *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1443
  P.position : 'a tlp -> ('a * Position.T) tlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1444
  P.arguments : Args.T list tlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1445
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1446
  val parse_src : Args.src tlp =
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1447
  P.position (P.xname -- P.arguments) >> Args.src ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1448
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1449
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1450
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1451
  val ((first_word, args), pos) = Args.dest_src asrc ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1452
  map Args.string_of args ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1453
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1454
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1455
  The \texttt{Args} structure contains more parsers and parser transformers 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1456
  for which the input source type is \texttt{Args.T list}.  For example,
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1457
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1458
  type 'a atlp = Args.T list -> 'a * Args.T list ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1459
  open Args ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1460
  nat : int atlp ; (* also Args.int *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1461
  thm_sel : PureThy.interval list atlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1462
  list : 'a atlp -> 'a list atlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1463
  attribs : (string -> string) -> Args.src list atlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1464
  opt_attribs : (string -> string) -> Args.src list atlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1465
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1466
  (* parse_atl_str : 'a atlp -> (string -> 'a * string) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1467
  given an Args.T list parser, to get a string parser *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1468
  fun parse_atl_str atlp str = 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1469
  let val (ats, rem_str) = parse_str P.arguments str ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1470
  val (res, rem_ats) = atlp ats ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1471
  in (res, String.concat (Library.separate " "
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1472
  (List.map Args.string_of rem_ats @ [rem_str]))) end ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1473
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1474
  parse_atl_str Args.int "-1-," ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1475
  parse_atl_str (Scan.option Args.int) "x1-," ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1476
  parse_atl_str Args.thm_sel "(1-,4,13-22)" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1477
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1478
  val (ats as atsrc :: _, "") = parse_atl_str (Args.attribs I)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1479
  "[THEN trans [THEN sym], simp, OF sym]" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1480
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1481
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1482
  From here, an attribute is interpreted using \texttt{Attrib.attribute}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1483
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1484
  \texttt{Args} has a large number of functions which parse an \texttt{Args.src}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1485
  and also refer to a generic context.  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1486
  Note the use of \texttt{Scan.lift} for this.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1487
  (as does \texttt{Attrib} - RETHINK THIS)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1488
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1489
  (\texttt{Args.syntax} shown below has type specialised)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1490
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1491
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1492
  type ('res, 'src) parse_fn = 'src -> 'res * 'src ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1493
  type 'a cgatlp = ('a, Context.generic * Args.T list) parse_fn ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1494
  Scan.lift : 'a atlp -> 'a cgatlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1495
  term : term cgatlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1496
  typ : typ cgatlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1497
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1498
  Args.syntax : string -> 'res cgatlp -> src -> ('res, Context.generic) parse_fn ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1499
  Attrib.thm : thm cgatlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1500
  Attrib.thms : thm list cgatlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1501
  Attrib.multi_thm : thm list cgatlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1502
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1503
  (* parse_cgatl_str : 'a cgatlp -> (string -> 'a * string) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1504
  given a (Context.generic * Args.T list) parser, to get a string parser *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1505
  fun parse_cgatl_str cgatlp str = 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1506
  let 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1507
    (* use the current generic context *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1508
    val generic = Context.Theory thy ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1509
    val (ats, rem_str) = parse_str P.arguments str ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1510
    (* ignore any change to the generic context *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1511
    val (res, (_, rem_ats)) = cgatlp (generic, ats) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1512
  in (res, String.concat (Library.separate " "
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1513
      (List.map Args.string_of rem_ats @ [rem_str]))) end ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1514
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1515
*}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1516
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1517
section{* Attributes, and the \texttt{Attrib} structure *}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1518
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1519
text {*
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1520
  The type \texttt{attribute} is declared in \texttt{src/Pure/thm.ML}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1521
  The source file for the \texttt{Attrib} structure is
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1522
  \texttt{src/Pure/Isar/attrib.ML}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1523
  Most attributes use a theorem to change a generic context (for example, 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1524
  by declaring that the theorem should be used, by default, in simplification),
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1525
  or change a theorem (which most often involves referring to the current
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1526
  theory). 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1527
  The functions \texttt{Thm.rule\_attribute} and
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1528
  \texttt{Thm.declaration\_attribute} create attributes of these kinds.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1529
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1530
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1531
  type attribute = Context.generic * thm -> Context.generic * thm;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1532
  type 'a trf = 'a -> 'a ; (* transformer of a given type *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1533
  Thm.rule_attribute  : (Context.generic -> thm -> thm) -> attribute ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1534
  Thm.declaration_attribute : (thm -> Context.generic trf) -> attribute ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1535
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1536
  Attrib.print_attributes : theory -> unit ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1537
  Attrib.pretty_attribs : Proof.context -> src list -> Pretty.T list ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1538
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1539
  List.app Pretty.writeln (Attrib.pretty_attribs ctxt ass) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1540
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1541
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1542
  An attribute is stored in a theory as indicated by:
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1543
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1544
  Attrib.add_attributes : 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1545
  (bstring * (src -> attribute) * string) list -> theory trf ; 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1546
  (*
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1547
  Attrib.add_attributes [("THEN", THEN_att, "resolution with rule")] ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1548
  *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1549
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1550
  where the first and third arguments are name and description of the attribute,
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1551
  and the second is a function which parses the attribute input text 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1552
  (including the attribute name, which has necessarily already been parsed).
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1553
  Here, \texttt{THEN\_att} is a function declared in the code for the
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1554
  structure \texttt{Attrib}, but not published in its signature.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1555
  The source file \texttt{src/Pure/Isar/attrib.ML} shows the use of 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1556
  \texttt{Attrib.add\_attributes} to add a number of attributes.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1557
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1558
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1559
  FullAttrib.THEN_att : src -> attribute ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1560
  FullAttrib.THEN_att atsrc (generic, ML_Context.thm "sym") ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1561
  FullAttrib.THEN_att atsrc (generic, ML_Context.thm "all_comm") ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1562
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1563
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1564
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1565
  Attrib.syntax : attribute cgatlp -> src -> attribute ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1566
  Attrib.no_args : attribute -> src -> attribute ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1567
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1568
  When this is called as \texttt{syntax scan src (gc, th)}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1569
  the generic context \texttt{gc} is used 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1570
  (and potentially changed to \texttt{gc'})
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1571
  by \texttt{scan} in parsing to obtain an attribute \texttt{attr} which would
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1572
  then be applied to \texttt{(gc', th)}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1573
  The source for parsing the attribute is the arguments part of \texttt{src},
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1574
  which must all be consumed by the parse.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1575
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1576
  For example, for \texttt{Attrib.no\_args attr src}, the attribute parser 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1577
  simply returns \texttt{attr}, requiring that the arguments part of
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1578
  \texttt{src} must be empty.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1579
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1580
  Some examples from \texttt{src/Pure/Isar/attrib.ML}, modified:
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1581
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1582
  fun rot_att_n n (gc, th) = (gc, rotate_prems n th) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1583
  rot_att_n : int -> attribute ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1584
  val rot_arg = Scan.lift (Scan.optional Args.int 1 : int atlp) : int cgatlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1585
  val rotated_att : src -> attribute =
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1586
  Attrib.syntax (rot_arg >> rot_att_n : attribute cgatlp) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1587
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1588
  val THEN_arg : int cgatlp = Scan.lift 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1589
  (Scan.optional (Args.bracks Args.nat : int atlp) 1 : int atlp) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1590
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1591
  Attrib.thm : thm cgatlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1592
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1593
  THEN_arg -- Attrib.thm : (int * thm) cgatlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1594
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1595
  fun THEN_att_n (n, tht) (gc, th) = (gc, th RSN (n, tht)) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1596
  THEN_att_n : int * thm -> attribute ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1597
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1598
  val THEN_att : src -> attribute = Attrib.syntax
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1599
  (THEN_arg -- Attrib.thm >> THEN_att_n : attribute cgatlp);
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1600
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1601
  The functions I've called \texttt{rot\_arg} and \texttt{THEN\_arg}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1602
  read an optional argument, which for \texttt{rotated} is an integer, 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1603
  and for \texttt{THEN} is a natural enclosed in square brackets;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1604
  the default, if the argument is absent, is 1 in each case.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1605
  Functions \texttt{rot\_att\_n} and \texttt{THEN\_att\_n} turn these into
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1606
  attributes, where \texttt{THEN\_att\_n} also requires a theorem, which is
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1607
  parsed by \texttt{Attrib.thm}.  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1608
  Infix operators \texttt{--} and \texttt{>>} are in the structure \texttt{Scan}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1609
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1610
*}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1611
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1612
section{* Methods, and the \texttt{Method} structure *}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1613
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1614
text {*
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1615
  The source file is \texttt{src/Pure/Isar/method.ML}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1616
  The type \texttt{method} is defined by the datatype declaration
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1617
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1618
  (* datatype method = Meth of thm list -> cases_tactic; *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1619
  RuleCases.NO_CASES : tactic -> cases_tactic ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1620
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1621
  In fact \texttt{RAW\_METHOD\_CASES} (below) is exactly the constructor
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1622
  \texttt{Meth}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1623
  A \texttt{cases\_tactic} is an elaborated version of a tactic.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1624
  \texttt{NO\_CASES tac} is a \texttt{cases\_tactic} which consists of a
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1625
  \texttt{cases\_tactic} without any further case information.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1626
  For further details see the description of structure \texttt{RuleCases} below.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1627
  The list of theorems to be passed to a method consists of the current
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1628
  \emph{facts} in the proof.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1629
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1630
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1631
  RAW_METHOD : (thm list -> tactic) -> method ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1632
  METHOD : (thm list -> tactic) -> method ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1633
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1634
  SIMPLE_METHOD : tactic -> method ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1635
  SIMPLE_METHOD' : (int -> tactic) -> method ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1636
  SIMPLE_METHOD'' : ((int -> tactic) -> tactic) -> (int -> tactic) -> method ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1637
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1638
  RAW_METHOD_CASES : (thm list -> cases_tactic) -> method ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1639
  METHOD_CASES : (thm list -> cases_tactic) -> method ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1640
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1641
  A method is, in its simplest form, a tactic; applying the method is to apply
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1642
  the tactic to the current goal state.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1643
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1644
  Applying \texttt{RAW\_METHOD tacf} creates a tactic by applying 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1645
  \texttt{tacf} to the current {facts}, and applying that tactic to the
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1646
  goal state.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1647
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1648
  \texttt{METHOD} is similar but also first applies
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1649
  \texttt{Goal.conjunction\_tac} to all subgoals.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1650
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1651
  \texttt{SIMPLE\_METHOD tac} inserts the facts into all subgoals and then
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1652
  applies \texttt{tacf}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1653
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1654
  \texttt{SIMPLE\_METHOD' tacf} inserts the facts and then
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1655
  applies \texttt{tacf} to subgoal 1.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1656
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1657
  \texttt{SIMPLE\_METHOD'' quant tacf} does this for subgoal(s) selected by
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1658
  \texttt{quant}, which may be, for example,
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1659
  \texttt{ALLGOALS} (all subgoals),
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1660
  \texttt{TRYALL} (try all subgoals, failure is OK),
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1661
  \texttt{FIRSTGOAL} (try subgoals until it succeeds once), 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1662
  \texttt{(fn tacf => tacf 4)} (subgoal 4), etc
16
5045dec52d2b polished
Christian Urban <urbanc@in.tum.de>
parents: 4
diff changeset
  1663
  (see the \texttt{Tactical} structure, FIXME) %%\cite[Chapter 4]{ref}).
4
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1664
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1665
  A method is stored in a theory as indicated by:
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1666
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1667
  Method.add_method : 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1668
  (bstring * (src -> Proof.context -> method) * string) -> theory trf ; 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1669
  ( *
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1670
  * )
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1671
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1672
  where the first and third arguments are name and description of the method,
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1673
  and the second is a function which parses the method input text 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1674
  (including the method name, which has necessarily already been parsed).
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1675
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1676
  Here, \texttt{xxx} is a function declared in the code for the
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1677
  structure \texttt{Method}, but not published in its signature.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1678
  The source file \texttt{src/Pure/Isar/method.ML} shows the use of 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1679
  \texttt{Method.add\_method} to add a number of methods.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1680
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1681
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1682
*}
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
  1683
(*>*)
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
  1684
end