CookBook/Parsing.thy
author Christian Urban <urbanc@in.tum.de>
Sun, 08 Feb 2009 08:45:25 +0000
changeset 104 5dcad9348e4d
parent 102 5e309df58557
child 105 f49dc7e96235
permissions -rw-r--r--
polished
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
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
     2
imports Base 
4
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
begin
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
     6
4
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
chapter {* Parsing *}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
text {*
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    11
  Isabelle distinguishes between \emph{outer} and \emph{inner} syntax. 
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    12
  Theory commands, such as \isacommand{definition}, \isacommand{inductive} and so
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    13
  on, belong to the outer syntax, whereas items inside double quotation marks, such 
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    14
  as terms, types and so on, belong to the inner syntax. For parsing inner syntax, 
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    15
  Isabelle uses a rather general and sophisticated algorithm due to Earley, which 
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    16
  is driven by priority grammars. Parsers for outer syntax are built up by functional
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    17
  parsing combinators. These combinators are a well-established technique for parsing, 
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
    18
  which has, for example, been described in Paulson's classic ML-book \cite{paulson-ml2}.
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
    19
  Isabelle developers are usually concerned with writing these outer syntax parsers, 
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
    20
  either for new definitional packages or for calling tactics with specific arguments. 
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
    21
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
    22
  \begin{readmore}
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
    23
  The library 
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
    24
  for writing parser combinators is split up, roughly, into two parts. 
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    25
  The first part consists of a collection of generic parser combinators defined
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    26
  in the structure @{ML_struct Scan} in the file 
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    27
  @{ML_file "Pure/General/scan.ML"}. The second part of the library consists of 
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    28
  combinators for dealing with specific token types, which are defined in the 
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    29
  structure @{ML_struct OuterParse} in the file 
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    30
  @{ML_file "Pure/Isar/outer_parse.ML"}.
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
    31
  \end{readmore}
38
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
*}
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    34
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
    35
section {* Building Generic Parsers *}
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    36
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    37
text {*
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    38
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    39
  Let us first have a look at parsing strings using generic parsing combinators. 
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
    40
  The function @{ML "(op $$)"} takes a string as argument and will ``consume'' this string from 
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    41
  a given input list of strings. ``Consume'' in this context means that it will 
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    42
  return a pair consisting of this string and the rest of the input list. 
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    43
  For example:
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    44
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
    45
  @{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
    46
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
    47
  @{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
    48
52
Christian Urban <urbanc@in.tum.de>
parents: 50
diff changeset
    49
  This function 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
    50
  @{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
    51
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
    52
  @{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
    53
                               "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
    54
  
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
    55
  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
    56
  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
    57
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    58
  \begin{itemize}
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
    59
  \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
    60
  might be explored. 
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
    61
  \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
    62
  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
    63
  \item @{text "ABORT"} is the exception that is raised when a dead end is reached. 
41
b11653b11bd3 further progress on the parsing section and tuning on the antiqu's
Christian Urban <urbanc@in.tum.de>
parents: 40
diff changeset
    64
  It is used for example in the function @{ML "(op !!)"} (see below).
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    65
  \end{itemize}
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    66
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
    67
  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
    68
  by the programmer (for example to handle them).
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
    69
  
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
    70
  Slightly more general than the parser @{ML "(op $$)"} is the function @{ML
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
    71
  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
    72
  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
    73
  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
    74
  [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
    75
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    76
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
    77
@{ML_response [display,gray] 
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
    78
"let 
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
    79
  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
    80
  val input1 = (explode \"hello\")
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    81
  val input2 = (explode \"world\")
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    82
in
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    83
    (hw input1, hw input2)
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    84
end"
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    85
    "((\"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
    86
52
Christian Urban <urbanc@in.tum.de>
parents: 50
diff changeset
    87
  Two parser can be connected in sequence by using the function @{ML "(op --)"}. 
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
    88
  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
    89
  sequence you can achieve by:
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
    90
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
    91
  @{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
    92
                          "(((\"h\", \"e\"), \"l\"), [\"l\", \"o\"])"}
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    93
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    94
  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
    95
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
    96
  If, as in the previous example, you want to parse a particular string, 
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
    97
  then one 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
    98
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
    99
  @{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
   100
                          "(\"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
   101
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   102
  Parsers that explore alternatives can be constructed using the function @{ML
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   103
  "(op ||)"}. 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
   104
  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
   105
  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
   106
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   107
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   108
@{ML_response [display,gray] 
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   109
"let 
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   110
  val hw = ($$ \"h\") || ($$ \"w\")
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   111
  val input1 = (explode \"hello\")
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   112
  val input2 = (explode \"world\")
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   113
in
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   114
  (hw input1, hw input2)
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   115
end"
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   116
  "((\"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
   117
52
Christian Urban <urbanc@in.tum.de>
parents: 50
diff changeset
   118
  The functions @{ML "(op |--)"} and @{ML "(op --|)"} work like the sequencing function 
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   119
  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
   120
  parser. For example:
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   121
  
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   122
@{ML_response [display,gray]
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   123
"let 
47
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_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
   125
  val just_h = ($$ \"h\") --| ($$ \"e\") 
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   126
  val input = (explode \"hello\")  
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   127
in 
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
   128
  (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
   129
end"
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   130
  "((\"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
   131
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   132
  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
   133
  @{text "p"}, if it succeeds; otherwise it returns 
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   134
  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
   135
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   136
@{ML_response [display,gray]
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   137
"let 
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   138
  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
   139
  val input1 = (explode \"hello\")
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   140
  val input2 = (explode \"world\")  
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   141
in 
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   142
  (p input1, p input2)
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   143
end" 
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   144
 "((\"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
   145
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   146
  The function @{ML Scan.option} works similarly, except no default value can
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   147
  be given. Instead, the result is wrapped as an @{text "option"}-type. For example:
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   148
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   149
@{ML_response [display,gray]
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   150
"let 
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   151
  val p = Scan.option ($$ \"h\")
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   152
  val input1 = (explode \"hello\")
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   153
  val input2 = (explode \"world\")  
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   154
in 
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   155
  (p input1, p input2)
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   156
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
   157
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   158
  The function @{ML "(op !!)"} 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
   159
  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
   160
  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
   161
  you might write:
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   162
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   163
  @{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
   164
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
   165
  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
   166
  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
   167
  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
   168
  @{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
   169
  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
   170
  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
   171
  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
   172
  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
   173
  failure of @{text r}, not by the absence of @{text q} in the input. This
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   174
  kind of situation can be avoided when using the function @{ML "(op !!)"}. 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   175
  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
   176
  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
   177
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   178
  
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   179
  @{ML [display,gray] "(!! (fn _ => \"foo\") ($$ \"h\"))"}
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   180
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
   181
  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
   182
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   183
  @{ML_response [display,gray] 
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   184
                "(!! (fn _ => \"foo\") ($$ \"h\")) (explode \"hello\")" 
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   185
                "(\"h\", [\"e\", \"l\", \"l\", \"o\"])"}
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   186
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
   187
  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
   188
  
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   189
  @{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
   190
                               "Exception ABORT raised"}
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   191
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   192
  then the parsing aborts and the error message @{text "foo"} is printed out. In order to
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   193
  see the error message properly, we 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
   194
  @{ML "Scan.error"}. For example:
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   195
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   196
  @{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
   197
                               "Exception Error \"foo\" raised"}
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   198
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   199
  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
   200
  (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
   201
  
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
   202
  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
   203
  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
   204
  then you have to write:
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   205
*}
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   206
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   207
ML{*fun p_followed_by_q p q r =
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   208
  let 
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   209
     val err_msg = (fn _ => p ^ " is not followed by " ^ q)
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   210
  in
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   211
    ($$ p -- (!! err_msg ($$ q))) || ($$ r -- $$ r)
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   212
  end *}
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   213
41
b11653b11bd3 further progress on the parsing section and tuning on the antiqu's
Christian Urban <urbanc@in.tum.de>
parents: 40
diff changeset
   214
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   215
text {*
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   216
  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
   217
  the input @{text [quotes] "holle"} 
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   218
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   219
  @{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
   220
                               "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
   221
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   222
  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
   223
 
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   224
  @{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
   225
                          "((\"w\", \"w\"), [\"o\", \"r\", \"l\", \"d\"])"}
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   226
  
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   227
  yields the expected parsing. 
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   228
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
   229
  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
   230
  often as it succeeds. For example:
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   231
  
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   232
  @{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
   233
                "([\"h\", \"h\", \"h\", \"h\"], [\"e\", \"l\", \"l\", \"o\"])"}
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   234
  
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   235
  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
   236
  @{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
   237
  succeeds at least once.
48
609f9ef73494 fixed FIXME's in fake responses
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
   238
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
   239
  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
   240
  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
   241
  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
   242
  them you can write:
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   243
  
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   244
  @{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
   245
                "([\"h\", \"h\", \"h\", \"h\"], [])"}
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   246
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   247
  @{ML Symbol.stopper} is the ``end-of-input'' indicator for parsing strings;
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
   248
  other stoppers need to be used when parsing tokens, for example. 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
   249
  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
   250
56
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   251
  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
   252
  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
   253
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   254
  @{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
   255
"let 
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 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
   257
   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
   258
in
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   259
   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
   260
end" 
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   261
"([\"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
   262
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   263
  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
   264
  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
   265
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   266
  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
   267
  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
   268
  
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   269
  @{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
   270
                               "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
   271
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   272
  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
   273
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   274
  @{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
   275
                          "(\"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
   276
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   277
  succeeds. 
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   278
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   279
  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
   280
  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
   281
  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
   282
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   283
  @{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
   284
"let 
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 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
   286
  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
   287
  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
   288
in
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 input1, 
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   290
   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
   291
end"
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
 ([\"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
   294
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
   295
  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
   296
  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
   297
  first the parser @{text p} and upon successful completion applies the 
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
   298
  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
   299
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   300
@{ML_response [display,gray]
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   301
"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
   302
  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
   303
in
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   304
  (($$ \"h\") -- ($$ \"e\") >> double) (explode \"hello\")
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   305
end"
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   306
"((\"hh\", \"ee\"), [\"l\", \"l\", \"o\"])"}
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   307
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   308
  doubles the two parsed input strings; or
59
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   309
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   310
  @{ML_response [display,gray] 
59
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   311
"let 
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   312
   val p = Scan.repeat (Scan.one Symbol.not_eof)
59
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   313
   val input = (explode \"foo bar foo\") 
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   314
in
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   315
   Scan.finite Symbol.stopper (p >> implode) input
59
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   316
end" 
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   317
"(\"foo bar foo\",[])"}
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   318
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
   319
  where the single-character strings in the parsed output are transformed
59
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   320
  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
   321
 
56
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   322
  \begin{exercise}\label{ex:scancmts}
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
   323
  Write a parser that parses an input string so that any comment enclosed
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
   324
  inside @{text "(*\<dots>*)"} is replaced by a the same comment but enclosed inside
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
   325
  @{text "(**\<dots>**)"} in the output string. To enclose a string, you can use the
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 56
diff changeset
   326
  function @{ML "enclose s1 s2 s" for s1 s2 s} which produces the string @{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
   327
  "s1 ^ s ^ s2" for s1 s2 s}.
56
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   328
  \end{exercise}
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   329
126646f2aa88 added a para on Scan.unless and an exercise about scanning comments
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   330
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   331
  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
   332
  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
   333
  untouched. For example
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   334
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   335
@{ML_response [display,gray]
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   336
"Scan.lift (($$ \"h\") -- ($$ \"e\")) (1,(explode \"hello\"))"
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   337
"((\"h\", \"e\"), (1, [\"l\", \"l\", \"o\"]))"}
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   338
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
   339
  (FIXME: In which situations is this useful? Give examples.) 
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   340
*}
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   341
41
b11653b11bd3 further progress on the parsing section and tuning on the antiqu's
Christian Urban <urbanc@in.tum.de>
parents: 40
diff changeset
   342
section {* Parsing Theory Syntax *}
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   343
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   344
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
   345
  Most of the time, however, Isabelle developers have to deal with parsing
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
   346
  tokens, not strings.  This is because the parsers for the theory syntax, as
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   347
  well as the parsers for the arguments of proof methods the type @{ML_type OuterLex.token} 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   348
  (which is identical to the type @{ML_type OuterParse.token}).  There are also handy 
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   349
  parsers for ML-expressions and ML-files.
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   350
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   351
  \begin{readmore}
40
35e1dff0d9bb more on the parsing section
Christian Urban <urbanc@in.tum.de>
parents: 39
diff changeset
   352
  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
   353
  @{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
   354
  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
   355
  \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
   356
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   357
  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
   358
  @{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
   359
  @{ML "Command" in OuterLex} for commands). Some token parsers take into account the 
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   360
  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
   361
*}  
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   362
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   363
text {*
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   364
  The first example shows how to generate a token list out of a string using
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   365
  the function @{ML "OuterSyntax.scan"}. It is given below @{ML "Position.none"}
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
   366
  as argument since, at the moment, we are not interested in generating
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
   367
  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
   368
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   369
@{ML_response_fake [display,gray] "OuterSyntax.scan Position.none \"hello world\"" 
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   370
"[Token (\<dots>,(Ident, \"hello\"),\<dots>), 
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   371
 Token (\<dots>,(Space, \" \"),\<dots>), 
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   372
 Token (\<dots>,(Ident, \"world\"),\<dots>)]"}
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   373
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   374
  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
   375
  @{text [quotes] "hello"} and @{text [quotes] "world"} do not match any
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   376
  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
   377
  the PolyML runtime system the result is printed as @{text [quotes] "?"}, instead of
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   378
  the tokens.} The second indicates a space.
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   379
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   380
  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
   381
  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
   382
  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
   383
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   384
@{ML_response_fake [display,gray]
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   385
"let
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   386
   val input = OuterSyntax.scan Position.none \"hello world\"
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   387
in
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   388
   filter OuterLex.is_proper input
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   389
end" 
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   390
"[Token (\<dots>,(Ident, \"hello\"), \<dots>), Token (\<dots>,(Ident, \"world\"), \<dots>)]"}
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   391
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
   392
  For convenience we define the function:
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   393
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   394
*}
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   395
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   396
ML{*fun filtered_input str = 
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   397
       filter OuterLex.is_proper (OuterSyntax.scan Position.none str) *}
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   398
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   399
text {*
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   400
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
   401
  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
   402
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   403
@{ML_response_fake [display,gray] 
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   404
"filtered_input \"inductive | for\"" 
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   405
"[Token (\<dots>,(Command, \"inductive\"),\<dots>), 
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   406
 Token (\<dots>,(Keyword, \"|\"),\<dots>), 
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   407
 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
   408
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
   409
  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
   410
  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
   411
  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
   412
  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
   413
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   414
@{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
   415
"let 
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
   416
  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
   417
in 
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
   418
  (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
   419
end" 
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
   420
"([\"}\",\"{\",\<dots>],[\"\<rightleftharpoons>\",\"\<leftharpoondown>\",\<dots>])"}
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   421
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
   422
  The parser @{ML "OuterParse.$$$"} parses a single keyword. For example:
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   423
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   424
@{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
   425
"let 
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   426
  val input1 = filtered_input \"where for\"
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   427
  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
   428
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
   429
  (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
   430
end"
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents: 43
diff changeset
   431
"((\"where\",\<dots>),(\"|\",\<dots>))"}
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents: 43
diff changeset
   432
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
   433
  Like before, you can sequentially connect parsers with @{ML "(op --)"}. 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
   434
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   435
@{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
   436
"let 
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   437
  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
   438
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
   439
  (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
   440
end"
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents: 43
diff changeset
   441
"((\"|\",\"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
   442
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   443
  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
   444
  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
   445
  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
   446
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   447
@{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
   448
"let 
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   449
  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
   450
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
   451
  (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
   452
end" 
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents: 43
diff changeset
   453
"([\"in\",\"in\",\"in\"],[\<dots>])"}
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents: 43
diff changeset
   454
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   455
  @{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
   456
  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
   457
  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
   458
  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
   459
  previous section, we can avoid this exception using the wrapper @{ML
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   460
  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
   461
  OuterLex.stopper}. We can write:
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   462
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   463
@{ML_response [display,gray]
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   464
"let 
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   465
  val input = filtered_input \"in | in | in\"
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   466
in 
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   467
  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
   468
         (OuterParse.enum \"|\" (OuterParse.$$$ \"in\")) input
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   469
end" 
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   470
"([\"in\",\"in\",\"in\"],[])"}
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   471
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   472
  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
   473
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   474
*}
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   475
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   476
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
   477
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   478
text {*
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   479
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   480
  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
   481
  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
   482
  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
   483
  with a relatively precise description of the failure. For example:
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   484
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   485
@{ML_response_fake [display,gray]
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   486
"let 
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   487
  val input = filtered_input \"in |\"
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   488
  val parse_bar_then_in = OuterParse.$$$ \"|\" -- OuterParse.$$$ \"in\"
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   489
in 
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   490
  parse (OuterParse.!!! parse_bar_then_in) input 
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   491
end"
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   492
"Exception ERROR \"Outer syntax error: keyword \"|\" expected, 
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   493
but keyword in was found\" raised"
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   494
}
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   495
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   496
  \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
   497
  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
   498
  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
   499
  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
   500
  \end{exercise}
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   501
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   502
  (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
   503
41
b11653b11bd3 further progress on the parsing section and tuning on the antiqu's
Christian Urban <urbanc@in.tum.de>
parents: 40
diff changeset
   504
*}
b11653b11bd3 further progress on the parsing section and tuning on the antiqu's
Christian Urban <urbanc@in.tum.de>
parents: 40
diff changeset
   505
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   506
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   507
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   508
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   509
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   510
section {* Positional Information *}
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   511
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   512
text {*
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   513
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   514
  @{ML OuterParse.position}
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   515
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   516
*}
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   517
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
   518
section {* Parsing Inner Syntax *}
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   519
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   520
ML{*let 
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
   521
  val input = OuterSyntax.scan Position.none "0"
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents: 43
diff changeset
   522
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
   523
  OuterParse.prop input
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   524
end*}
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   525
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   526
text {* (FIXME funny output for a proposition) *}
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   527
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   528
section {* Parsing Specifications *}
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   529
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   530
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
   531
  There are a number of special purpose parsers that help with parsing specifications
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
   532
  of functions, inductive definitions and so on. For example the function 
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
   533
  @{ML OuterParse.target} reads a target in order to indicate a locale.
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   534
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   535
@{ML_response [display,gray]
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   536
"let 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   537
  val input = filtered_input \"(in test)\"
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   538
in 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   539
  parse OuterParse.target input 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   540
end" "(\"test\",[])"}
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   541
  
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
   542
  The function @{ML OuterParse.opt_target} makes this parser ``optional'', that
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
   543
  is wrapping the result into an option type and returning @{ML NONE} if no
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
   544
  target is present.
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   545
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   546
  The function @{ML OuterParse.fixes} reads a list of constants
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
   547
  that can include type annotations and syntax translations.
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
   548
  The list is separated by \isacommand{and}. For example:
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   549
*}
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   550
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   551
text {*
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   552
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   553
@{ML_response [display,gray]
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   554
"let
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   555
  val input = filtered_input 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   556
        \"foo::\\\"int \<Rightarrow> bool\\\" (\\\"FOO\\\" [100] 100) and bar::nat and blonk\"
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   557
in
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   558
   parse OuterParse.fixes input
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   559
end"
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   560
"([(foo, SOME \<dots>, Mixfix (\"FOO\",[100],100)), 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   561
  (bar, SOME \<dots>, NoSyn), 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   562
  (blonk, NONE, NoSyn)],[])"}  
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   563
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
   564
  Whenever types are given, then they are stored in the @{ML SOME}s. 
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
   565
  Note that we had to write @{text "\\\"int \<Rightarrow> bool\\\""} to properly escape 
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
   566
  the double quotes in the compound type.
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   567
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   568
  @{ML OuterParse.for_fixes} is an ``optional'' that prefixes 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   569
  @{ML OuterParse.fixes} with the comman \isacommand{for}.
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
   570
  (FIXME give an example and explain more)
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
   571
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
   572
@{ML_response [display,gray]
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
   573
"let
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
   574
  val input = filtered_input 
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
   575
        \"for foo::\\\"int \<Rightarrow> bool\\\" (\\\"FOO\\\" [100] 100) and bar::nat and blonk\"
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
   576
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
   577
   parse OuterParse.for_fixes input
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
   578
end"
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
   579
"([(foo, SOME \<dots>, Mixfix (\"FOO\",[100],100)), 
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
   580
  (bar, SOME \<dots>, NoSyn), 
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
   581
  (blonk, NONE, NoSyn)],[])"}  
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
   582
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   583
*}
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   584
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
   585
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   586
ML{*let 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   587
  val input = filtered_input "test_lemma[intro,foo,elim,dest!,bar]:"
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   588
in 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   589
  parse (SpecParse.thm_name ":") input 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   590
     |> fst |> snd |> (Attrib.pretty_attribs @{context}) |> (map Pretty.string_of)
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   591
end*}
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   592
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   593
text {*
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   594
  (FIXME: why is intro, elim and dest treated differently from bar?) 
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   595
*}
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   596
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   597
text {*
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   598
@{ML_response [display,gray]
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   599
"let
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   600
  val input = filtered_input
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   601
     (\"even and odd \" ^  
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   602
      \"where \" ^
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   603
      \"  even0[intro]: \\\"even 0\\\" \" ^ 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   604
      \"| evenS[intro]: \\\"odd n \<Longrightarrow> even (Suc n)\\\" \" ^ 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   605
      \"| oddS[intro]:  \\\"even n \<Longrightarrow> odd (Suc n)\\\"\")
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   606
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   607
  val parser = 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   608
    OuterParse.opt_target --
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   609
    OuterParse.fixes -- 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   610
    OuterParse.for_fixes --
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   611
    Scan.optional 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   612
        (OuterParse.$$$ \"where\" |--
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   613
           OuterParse.!!! 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   614
             (OuterParse.enum1 \"|\" 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   615
                (SpecParse.opt_thm_name \":\" -- OuterParse.prop))) []
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   616
in
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   617
  parse parser input
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   618
end"
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   619
"((((NONE, [(even, NONE, NoSyn), (odd, NONE, NoSyn)]), []),
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   620
     [((even0,\<dots>), \"\\^E\\^Ftoken\\^Eeven 0\\^E\\^F\\^E\"),
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   621
      ((evenS,\<dots>), \"\\^E\\^Ftoken\\^Eodd n \<Longrightarrow> even (Suc n)\\^E\\^F\\^E\"),
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   622
      ((oddS,\<dots>), \"\\^E\\^Ftoken\\^Eeven n \<Longrightarrow> odd (Suc n)\\^E\\^F\\^E\")]), [])"}
50
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   623
*}
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   624
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   625
text {* 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   626
  The (outer?) parser for the package: returns optionally a locale; 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   627
  a list of predicate constants with optional type-annotation and 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   628
  optional syntax-annotation; a list of for-fixes (fixed parameters); 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   629
  and a list of rules where each rule has optionally a name and an attribute.
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   630
*}
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   631
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
   632
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
   633
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   634
text {*
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   635
  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
   636
  need to be implemented. While this is not difficult on the ML-level,
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   637
  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
   638
  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
   639
  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
   640
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   641
  To keep things simple, let us start with a ``silly'' command that does nothing 
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   642
  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
   643
  defined as:
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   644
*}
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   645
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   646
ML{*let
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   647
  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
   648
  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
   649
in
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   650
  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
   651
end *}
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   652
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   653
text {*
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   654
  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
   655
  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
   656
  parser producing a top-level transition function (its purpose will also explained
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   657
  later). 
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   658
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   659
  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
   660
  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
   661
  recognise \isacommand{foobar} as a command. Such a keyword file can be
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   662
  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
   663
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   664
  @{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
   665
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   666
  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
   667
  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
   668
  "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
   669
  present (in order to extract the keywords from them). To generate these log
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   670
  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
   671
  @{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
   672
  complete code.
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   673
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   674
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   675
  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   676
  \begin{figure}[t]
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   677
  \begin{graybox}\small
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   678
  \isacommand{theory}~@{text Command}\\
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   679
  \isacommand{imports}~@{text Main}\\
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   680
  \isacommand{begin}\\
85
b02904872d6b better handling of {* and *}
Christian Urban <urbanc@in.tum.de>
parents: 81
diff changeset
   681
  \isacommand{ML}~@{text "\<verbopen>"}\\
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   682
  @{ML
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   683
"let
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   684
  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
   685
  val kind = OuterKeyword.thy_decl
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   686
in
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   687
  OuterSyntax.command \"foobar\" \"description of foobar\" kind do_nothing
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   688
end"}\\
85
b02904872d6b better handling of {* and *}
Christian Urban <urbanc@in.tum.de>
parents: 81
diff changeset
   689
  @{text "\<verbclose>"}\\
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   690
  \isacommand{end}
80
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   691
  \end{graybox}
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   692
  \caption{\small The file @{text "Command.thy"} is necessary for generating a log 
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   693
  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
   694
  the command \isacommand{foobar}.\label{fig:commandtheory}}
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   695
  \end{figure}
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   696
  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   697
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   698
  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
   699
  @{text "Pure"}, @{text "HOL"} and @{text "Pure-ProofGeneral"}, as well as
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   700
  the log file for the theory @{text "Command.thy"}, which contains the new
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   701
  \isacommand{foobar}-command. If you target other logics besides HOL, such
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   702
  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
   703
  
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   704
  @{text Pure} and @{text HOL} are usually compiled during the installation of
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   705
  Isabelle. So log files for them should be already available. If not, then
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   706
  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
   707
  distribution.
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   708
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   709
  @{text [display] 
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   710
"$ ./build -m \"Pure\"
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   711
$ ./build -m \"HOL\""}
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   712
  
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
   713
  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
   714
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   715
  @{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
   716
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   717
  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
   718
  with:
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   719
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   720
  @{text [display] "$ isabelle mkdir FoobarCommand"}
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   721
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
   722
  This generates a directory containing the files: 
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   723
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   724
  @{text [display] 
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   725
"./IsaMakefile
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   726
./FoobarCommand/ROOT.ML
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   727
./FoobarCommand/document
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   728
./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
   729
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   730
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   731
  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
   732
  and add the line 
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   733
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   734
  @{text [display] "use_thy \"Command\";"} 
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   735
  
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
   736
  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
   737
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   738
  @{text [display] "$ isabelle make"}
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   739
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   740
  If the compilation succeeds, you have finally created all the necessary log files. 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   741
  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
   742
  
66
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   743
  @{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
   744
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   745
  or something similar depending on your Isabelle distribution and architecture.
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   746
  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
   747
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   748
  @{text [display] "$ ISABELLE_LOGS=\"$(isabelle getenv -b ISABELLE_OUTPUT)\"/log"}
Christian Urban <urbanc@in.tum.de>
parents: 65
diff changeset
   749
 
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
   750
  on the Unix prompt. The 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
   751
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   752
  @{text [display] 
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   753
"Pure.gz
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   754
HOL.gz
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   755
Pure-ProofGeneral.gz
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   756
HOL-FoobarCommand.gz"} 
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   757
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   758
  From them you can create the keyword files. Assuming the name 
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   759
  of the directory is in  @{text "$ISABELLE_LOGS"},
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   760
  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
   761
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   762
@{text [display]
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   763
"$ isabelle keywords -k foobar 
80
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   764
   $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
   765
80
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   766
  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
   767
  the string @{text "foobar"} twice.\footnote{To see whether things are fine, check
80
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   768
  that @{text "grep foobar"} on this file returns something
86
fdb9ea86c2a3 polished
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
   769
  non-empty.}  This keyword file needs to
80
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   770
  be copied into the directory @{text "~/.isabelle/etc"}. To make Isabelle
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   771
  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
   772
  "-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
   773
80
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   774
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
   775
  @{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
   776
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   777
  If you now build a theory on top of @{text "Command.thy"}, 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   778
  then the command \isacommand{foobar} can be used. 
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   779
  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
   780
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   781
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   782
  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
   783
  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
   784
  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
   785
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   786
  The crucial part of a command is the function that determines the behaviour
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   787
  of the command. In the code above we used a ``do-nothing''-function, which
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   788
  because of @{ML Scan.succeed} does not parse any argument, but immediately
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   789
  returns the simple toplevel function @{ML "Toplevel.theory I"}. We can
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   790
  replace this code by a function that first parses a proposition (using the
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   791
  parser @{ML OuterParse.prop}), then prints out the tracing
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   792
  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
   793
  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
   794
*}
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   795
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   796
ML{*let
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   797
  fun trace_top_lvl str = 
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   798
     Toplevel.theory (fn thy => (tracing str; thy))
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   799
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   800
  val trace_prop = OuterParse.prop >> trace_top_lvl
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   801
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   802
  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
   803
in
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   804
  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
   805
end *}
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   806
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   807
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
   808
  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
   809
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   810
  \begin{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   811
  \isacommand{foobar}~@{text [quotes] "True \<and> False"}\\
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   812
  @{text "> \"True \<and> False\""}
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   813
  \end{isabelle}
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   814
  
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   815
  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
   816
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   817
  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
   818
  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
   819
  arguments are processed. Examples of this kind of commands are
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   820
  \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
   821
  commands are expected to parse some arguments, for example a proposition,
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   822
  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
   823
  \isacommand{lemma}) or prove some other properties (for example
104
5dcad9348e4d polished
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   824
  \isacommand{function}). To achieve this kind of behaviour, you have to use the kind
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   825
  indicator @{ML thy_goal in OuterKeyword}.
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   826
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   827
  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
   828
  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
   829
  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
   830
*}
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   831
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   832
ML%linenumbers{*let
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   833
  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
   834
    let
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   835
      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
   836
    in
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   837
      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
   838
    end;
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   839
  
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   840
  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
   841
      (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
   842
                    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
   843
  
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   844
  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
   845
in
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   846
  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
   847
end *}
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   848
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   849
text {*
86
fdb9ea86c2a3 polished
Christian Urban <urbanc@in.tum.de>
parents: 85
diff changeset
   850
  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
   851
  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
   852
  @{ML Syntax.read_prop}, which converts a string into a proper proposition.
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   853
  In Line 6 the function @{ML Proof.theorem_i} starts the proof for the
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   854
  proposition. Its argument @{ML NONE} stands for a locale (which we chose to
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   855
  omit); the argument @{ML "(K I)"} stands for a function that determines what
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   856
  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
   857
  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
   858
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
   859
  (FIXME: explain @{ML Toplevel.print} etc)
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   860
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
   861
  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
   862
  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
   863
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   864
  \begin{isabelle}
80
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   865
  \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
   866
  @{text "goal (1 subgoal):"}\\
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   867
  @{text "1. True \<and> True"}
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   868
  \end{isabelle}
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   869
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
   870
  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
   871
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   872
  \begin{isabelle}
80
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   873
  \isacommand{foobar}~@{text [quotes] "True \<and> True"}\\
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   874
  \isacommand{apply}@{text "(rule conjI)"}\\
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   875
  \isacommand{apply}@{text "(rule TrueI)+"}\\
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   876
  \isacommand{done}
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   877
  \end{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
   878
68
e7519207c2b7 added more to the "new command section" and tuning
Christian Urban <urbanc@in.tum.de>
parents: 67
diff changeset
   879
  
101
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   880
  (FIXME What do @{ML "Toplevel.theory"} 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   881
  @{ML "Toplevel.print"} 
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   882
  @{ML Toplevel.local_theory}?)
Christian Urban <urbanc@in.tum.de>
parents: 86
diff changeset
   883
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   884
  (FIXME read a name and show how to store theorems)
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   885
65
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   886
*}
c8e9a4f97916 tuned and added a section about creating keyword files
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
   887
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   888
(*<*)
80
Christian Urban <urbanc@in.tum.de>
parents: 75
diff changeset
   889
38
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   890
chapter {* Parsing *}
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   891
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   892
text {*
e21b2f888fa2 added a preliminary section about parsing
Christian Urban <urbanc@in.tum.de>
parents: 16
diff changeset
   893
4
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   894
  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
   895
  including:
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   896
  \begin{itemize}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   897
  \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
   898
  or simplified versions of such code
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   899
  \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
   900
  (or specialisations of their types)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   901
  \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
   902
  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
   903
  \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
   904
  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
   905
  \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
   906
  \end{itemize}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   907
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   908
*}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   909
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   910
section {* Parsing Isar input *}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   911
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   912
text {*
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   913
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   914
  The typical parsing function has the type
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   915
  \texttt{'src -> 'res * 'src}, with input  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   916
  of type \texttt{'src}, returning a result 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   917
  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
   918
  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
   919
  (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
   920
  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
   921
  value of type \texttt{'res}). 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   922
  An exception is raised if an appropriate value 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   923
  cannot be produced from the input.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   924
  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
   925
  for the failure of a parse.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   926
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   927
  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
   928
  which is of type 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   929
  \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
   930
  (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
   931
  However, much of the discussion at 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   932
  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
   933
  is relevant.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   934
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   935
  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
   936
  as follows:
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   937
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   938
  open StringCvt ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   939
  type ('res, 'src) ex_reader = 'src -> 'res * 'src
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   940
  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
   941
  fun ex_reader rdr src = Option.valOf (rdr src) ;
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
   942
  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
   943
  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
   944
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   945
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   946
*}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   947
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   948
section{* The \texttt{Scan} structure *}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   949
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   950
text {* 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   951
  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
   952
  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
   953
  of the type \texttt{'src -> 'res * 'src}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   954
  Three exceptions are used:
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   955
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   956
  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
   957
  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
   958
  exception ABORT of string;        (*dead end*)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   959
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   960
  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
   961
  symbols) are declared as infix.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   962
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   963
  Some functions from that structure are
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   964
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   965
  |-- : ('src -> 'res1 * 'src') * ('src' -> 'res2 * 'src'') ->
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   966
  'src -> 'res2 * 'src''
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   967
  --| : ('src -> 'res1 * 'src') * ('src' -> 'res2 * 'src'') ->
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   968
  'src -> 'res1 * 'src''
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   969
  -- : ('src -> 'res1 * 'src') * ('src' -> 'res2 * 'src'') ->
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   970
  'src -> ('res1 * 'res2) * 'src''
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   971
  ^^ : ('src -> string * 'src') * ('src' -> string * 'src'') ->
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   972
  'src -> string * 'src''
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   973
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   974
  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
   975
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   976
  \texttt{|--} and \texttt{--|} 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   977
  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
   978
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   979
  \texttt{--} returns both.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   980
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   981
  \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
   982
  (which must be strings).
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   983
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   984
  Note how, although the types 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   985
  \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
   986
  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
   987
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   988
  :-- : ('src -> 'res1 * 'src') * ('res1 -> 'src' -> 'res2 * 'src'') ->
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   989
  'src -> ('res1 * 'res2) * 'src''
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   990
  :|-- : ('src -> 'res1 * 'src') * ('res1 -> 'src' -> 'res2 * 'src'') ->
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   991
  'src -> 'res2 * 'src''
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   992
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   993
  These are similar to \texttt{|--} and \texttt{--|},
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   994
  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
   995
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   996
  >> : ('src -> 'res1 * 'src') * ('res1 -> 'res2) -> 'src -> 'res2 * 'src'
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   997
  || : ('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
   998
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   999
  \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
  1000
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1001
  \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
  1002
  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
  1003
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1004
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1005
  succeed : 'res -> ('src -> 'res * 'src) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1006
  fail : ('src -> 'res_src) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1007
  !! : ('src * string option -> string) -> 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1008
  ('src -> 'res_src) -> ('src -> 'res_src) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1009
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1010
  \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
  1011
  \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
  1012
  \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
  1013
  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
  1014
  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
  1015
  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
  1016
  it won't recover by trying \texttt{parse2}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1017
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1018
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1019
  one : ('si -> bool) -> ('si list -> 'si * 'si list) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1020
  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
  1021
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1022
  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
  1023
  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
  1024
  On other failures they raise \texttt{FAIL NONE} 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1025
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1026
  \texttt{one p} takes the first
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1027
  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
  1028
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1029
  \texttt{some f} takes the first
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1030
  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
  1031
  \texttt{NONE}.  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1032
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1033
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1034
  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
  1035
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1036
  \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
  1037
  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
  1038
  it fails, raising \texttt{MORE NONE}.
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
  \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
  1041
  does not satisfy \texttt{p}.  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1042
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1043
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1044
  option : ('src -> 'res * 'src) -> ('src -> 'res option * 'src)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1045
  optional : ('src -> 'res * 'src) -> 'res -> ('src -> 'res * 'src)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1046
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1047
  \texttt{option}: 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1048
  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
  1049
  or raises \texttt{FAIL \_},
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1050
  \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
  1051
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1052
  \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
  1053
  \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
  1054
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1055
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1056
  repeat : ('src -> 'res * 'src) -> 'src -> 'res list * 'src
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1057
  repeat1 : ('src -> 'res * 'src) -> 'src -> 'res list * 'src
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1058
  bulk : ('src -> 'res * 'src) -> 'src -> 'res list * 'src 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1059
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1060
  \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
  1061
  \texttt{f} fails with \texttt{FAIL \_}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1062
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1063
  \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
  1064
  successful parse.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1065
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1066
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1067
  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
  1068
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1069
  \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
  1070
  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
  1071
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1072
  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
  1073
  HOW DO THEY WORK ?? TO BE COMPLETED
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1074
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1075
  dest_lexicon: lexicon -> string list ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1076
  make_lexicon: string list list -> lexicon ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1077
  empty_lexicon: lexicon ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1078
  extend_lexicon: string list list -> lexicon -> lexicon ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1079
  merge_lexicons: lexicon -> lexicon -> lexicon ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1080
  is_literal: lexicon -> string list -> bool ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1081
  literal: lexicon -> string list -> string list * string list ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1082
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1083
  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
  1084
  by:
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1085
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1086
  val (command_lexicon, keyword_lexicon) = OuterSyntax.get_lexicons () ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1087
  val commands = Scan.dest_lexicon command_lexicon ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1088
  val keywords = Scan.dest_lexicon keyword_lexicon ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1089
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1090
*}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1091
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1092
section{* The \texttt{OuterLex} structure *}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1093
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1094
text {*
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1095
  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
  1096
  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
  1097
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1098
  structure T = OuterLex;
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
  This structure defines the type \texttt{token}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1101
  (The types
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1102
  \texttt{OuterLex.token},
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1103
  \texttt{OuterParse.token} and
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1104
  \texttt{SpecParse.token} are all the same).
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1105
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1106
  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
  1107
  functions is \texttt{token list}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1108
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1109
  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
  1110
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1111
  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
  1112
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1113
  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
  1114
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1115
*}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1116
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1117
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1118
4
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1119
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
  1120
ML{*
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1121
  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
  1122
   "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
  1123
*}
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1124
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
  1125
ML{*
4
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1126
  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
  1127
*}
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1128
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
  1129
ML{*
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1130
  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
  1131
*}
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1132
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
  1133
ML{*
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1134
  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
  1135
*}  
4
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1136
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
  1137
ML{*
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1138
  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
  1139
*}
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1140
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
  1141
ML{*
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1142
  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
  1143
*}
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1144
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
  1145
ML{*
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
  1146
  OuterLex.stopper
4
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1147
*}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1148
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1149
text {*
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1150
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1151
  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
  1152
  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
  1153
  comments to have been filtered out.
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
  There is a special end-of-file token:
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1156
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1157
  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
  1158
  (* end of file token *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1159
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1160
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1161
*}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1162
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1163
section {* The \texttt{OuterParse} structure *}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1164
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1165
text {*
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1166
  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
  1167
  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
  1168
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1169
  structure P = OuterParse;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1170
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1171
  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
  1172
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1173
  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
  1174
  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
  1175
  \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
  1176
  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
  1177
  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
  1178
  (eg, \texttt{name, xname, text, typ}).
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1179
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1180
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1181
  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
  1182
  $$$ : string -> string tlp
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1183
  nat : int tlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1184
  maybe : 'a tlp -> 'a option tlp ;
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
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1187
  \texttt{\$\$\$ s} returns the first token,
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1188
  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
  1189
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1190
  \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
  1191
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1192
  \texttt{maybe}: if \texttt{p} returns \texttt{r}, 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1193
  then \texttt{maybe p} returns \texttt{SOME r} ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1194
  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
  1195
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1196
  A few examples:
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1197
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1198
  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
  1199
  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
  1200
  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
  1201
  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
  1202
  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
  1203
  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
  1204
  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
  1205
  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
  1206
  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
  1207
  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
  1208
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1209
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1210
  The following code helps run examples:
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
  fun parse_str tlp str = 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1213
  let val toks : token list = OuterSyntax.scan str ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1214
  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
  1215
  val (res, rem_toks) = tlp proper_toks ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1216
  val rem_str = String.concat
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1217
  (Library.separate " " (List.map T.unparse rem_toks)) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1218
  in (res, rem_str) end ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1219
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1220
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1221
  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
  1222
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1223
  val type_args =
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1224
  type_ident >> Library.single ||
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1225
  $$$ "(" |-- !!! (list1 type_ident --| $$$ ")") ||
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1226
  Scan.succeed [];
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1227
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1228
  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
  1229
  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
  1230
  list.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1231
  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
  1232
  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
  1233
  and then a ")" which is also ignored.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1234
  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
  1235
  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
  1236
  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
  1237
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1238
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1239
  fun triple2 (x, (y, z)) = (x, y, z);
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1240
  val arity = xname -- ($$$ "::" |-- !!! (
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1241
  Scan.optional ($$$ "(" |-- !!! (list1 sort --| $$$ ")")) []
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1242
  -- sort)) >> triple2;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1243
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1244
  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
  1245
  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
  1246
  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
  1247
  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
  1248
  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
  1249
  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
  1250
  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
  1251
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1252
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1253
  parse_str P.type_args "('a, 'b) ntyp" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1254
  parse_str P.type_args "'a ntyp" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1255
  parse_str P.type_args "ntyp" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1256
  parse_str P.arity "ty :: tycl" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1257
  parse_str P.arity "ty :: (tycl1, tycl2) tycl" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1258
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1259
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1260
*}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1261
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1262
section {* The \texttt{SpecParse} structure *}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1263
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1264
text {*
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1265
  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
  1266
  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
  1267
  For example, 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1268
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1269
  open SpecParse ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1270
  attrib : Attrib.src tok_rdr ; 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1271
  attribs : Attrib.src list tok_rdr ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1272
  opt_attribs : Attrib.src list tok_rdr ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1273
  xthm : (thmref * Attrib.src list) tok_rdr ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1274
  xthms1 : (thmref * Attrib.src list) list tok_rdr ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1275
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1276
  parse_str attrib "simp" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1277
  parse_str opt_attribs "hello" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1278
  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
  1279
  map Args.dest_src ass ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1280
  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
  1281
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1282
  parse_str xthm "mythm [attr]" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1283
  parse_str xthms1 "thm1 [attr] thms2" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1284
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1285
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1286
  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
  1287
  structure, described below.
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
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1290
section{* The \texttt{Args} structure *}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1291
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1292
text {*
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1293
  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
  1294
  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
  1295
  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
  1296
  \texttt{Args.src} and \texttt{Args.dest\_src}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1297
  are in fact the constructor and destructor functions.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1298
  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
  1299
  are in fact \texttt{Args.src}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1300
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1301
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1302
  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
  1303
  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
  1304
  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
  1305
  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
  1306
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1307
  val thy = ML_Context.the_context () ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1308
  val ctxt = ProofContext.init thy ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1309
  map (pr_src ctxt) ass ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1310
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1311
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1312
  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
  1313
  ``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
  1314
  input.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1315
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1316
  (* how an Args.src is parsed *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1317
  P.position : 'a tlp -> ('a * Position.T) tlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1318
  P.arguments : Args.T list tlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1319
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1320
  val parse_src : Args.src tlp =
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1321
  P.position (P.xname -- P.arguments) >> Args.src ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1322
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1323
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1324
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1325
  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
  1326
  map Args.string_of args ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1327
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1328
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1329
  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
  1330
  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
  1331
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1332
  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
  1333
  open Args ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1334
  nat : int atlp ; (* also Args.int *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1335
  thm_sel : PureThy.interval list atlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1336
  list : 'a atlp -> 'a list atlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1337
  attribs : (string -> string) -> Args.src list atlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1338
  opt_attribs : (string -> string) -> Args.src list atlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1339
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1340
  (* parse_atl_str : 'a atlp -> (string -> 'a * string) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1341
  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
  1342
  fun parse_atl_str atlp str = 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1343
  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
  1344
  val (res, rem_ats) = atlp ats ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1345
  in (res, String.concat (Library.separate " "
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1346
  (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
  1347
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1348
  parse_atl_str Args.int "-1-," ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1349
  parse_atl_str (Scan.option Args.int) "x1-," ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1350
  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
  1351
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1352
  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
  1353
  "[THEN trans [THEN sym], simp, OF sym]" ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1354
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1355
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1356
  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
  1357
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1358
  \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
  1359
  and also refer to a generic context.  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1360
  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
  1361
  (as does \texttt{Attrib} - RETHINK THIS)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1362
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1363
  (\texttt{Args.syntax} shown below has type specialised)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1364
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1365
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1366
  type ('res, 'src) parse_fn = 'src -> 'res * 'src ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1367
  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
  1368
  Scan.lift : 'a atlp -> 'a cgatlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1369
  term : term cgatlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1370
  typ : typ cgatlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1371
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1372
  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
  1373
  Attrib.thm : thm cgatlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1374
  Attrib.thms : thm list cgatlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1375
  Attrib.multi_thm : thm list cgatlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1376
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1377
  (* parse_cgatl_str : 'a cgatlp -> (string -> 'a * string) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1378
  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
  1379
  fun parse_cgatl_str cgatlp str = 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1380
  let 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1381
    (* use the current generic context *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1382
    val generic = Context.Theory thy ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1383
    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
  1384
    (* ignore any change to the generic context *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1385
    val (res, (_, rem_ats)) = cgatlp (generic, ats) ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1386
  in (res, String.concat (Library.separate " "
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1387
      (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
  1388
  \end{verbatim}
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
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1391
section{* Attributes, and the \texttt{Attrib} structure *}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1392
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1393
text {*
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1394
  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
  1395
  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
  1396
  \texttt{src/Pure/Isar/attrib.ML}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1397
  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
  1398
  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
  1399
  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
  1400
  theory). 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1401
  The functions \texttt{Thm.rule\_attribute} and
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1402
  \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
  1403
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1404
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1405
  type attribute = Context.generic * thm -> Context.generic * thm;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1406
  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
  1407
  Thm.rule_attribute  : (Context.generic -> thm -> thm) -> attribute ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1408
  Thm.declaration_attribute : (thm -> Context.generic trf) -> attribute ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1409
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1410
  Attrib.print_attributes : theory -> unit ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1411
  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
  1412
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1413
  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
  1414
  \end{verbatim}
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
  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
  1417
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1418
  Attrib.add_attributes : 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1419
  (bstring * (src -> attribute) * string) list -> theory trf ; 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1420
  (*
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1421
  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
  1422
  *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1423
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1424
  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
  1425
  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
  1426
  (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
  1427
  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
  1428
  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
  1429
  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
  1430
  \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
  1431
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1432
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1433
  FullAttrib.THEN_att : src -> attribute ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1434
  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
  1435
  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
  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
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1439
  Attrib.syntax : attribute cgatlp -> src -> attribute ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1440
  Attrib.no_args : attribute -> src -> attribute ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1441
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1442
  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
  1443
  the generic context \texttt{gc} is used 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1444
  (and potentially changed to \texttt{gc'})
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1445
  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
  1446
  then be applied to \texttt{(gc', th)}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1447
  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
  1448
  which must all be consumed by the parse.
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
  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
  1451
  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
  1452
  \texttt{src} must be empty.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1453
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1454
  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
  1455
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1456
  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
  1457
  rot_att_n : int -> attribute ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1458
  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
  1459
  val rotated_att : src -> attribute =
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1460
  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
  1461
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1462
  val THEN_arg : int cgatlp = Scan.lift 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1463
  (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
  1464
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1465
  Attrib.thm : thm cgatlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1466
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1467
  THEN_arg -- Attrib.thm : (int * thm) cgatlp ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1468
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1469
  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
  1470
  THEN_att_n : int * thm -> attribute ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1471
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1472
  val THEN_att : src -> attribute = Attrib.syntax
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1473
  (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
  1474
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1475
  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
  1476
  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
  1477
  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
  1478
  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
  1479
  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
  1480
  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
  1481
  parsed by \texttt{Attrib.thm}.  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1482
  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
  1483
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1484
*}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1485
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1486
section{* Methods, and the \texttt{Method} structure *}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1487
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1488
text {*
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1489
  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
  1490
  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
  1491
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1492
  (* datatype method = Meth of thm list -> cases_tactic; *)
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1493
  RuleCases.NO_CASES : tactic -> cases_tactic ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1494
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1495
  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
  1496
  \texttt{Meth}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1497
  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
  1498
  \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
  1499
  \texttt{cases\_tactic} without any further case information.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1500
  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
  1501
  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
  1502
  \emph{facts} in the proof.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1503
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1504
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1505
  RAW_METHOD : (thm list -> tactic) -> method ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1506
  METHOD : (thm list -> tactic) -> method ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1507
  
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1508
  SIMPLE_METHOD : tactic -> method ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1509
  SIMPLE_METHOD' : (int -> tactic) -> method ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1510
  SIMPLE_METHOD'' : ((int -> tactic) -> tactic) -> (int -> tactic) -> method ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1511
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1512
  RAW_METHOD_CASES : (thm list -> cases_tactic) -> method ;
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1513
  METHOD_CASES : (thm list -> cases_tactic) -> method ;
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
  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
  1516
  the tactic to the current goal state.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1517
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1518
  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
  1519
  \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
  1520
  goal state.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1521
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1522
  \texttt{METHOD} is similar but also first applies
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1523
  \texttt{Goal.conjunction\_tac} to all subgoals.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1524
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1525
  \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
  1526
  applies \texttt{tacf}.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1527
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1528
  \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
  1529
  applies \texttt{tacf} to subgoal 1.
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1530
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1531
  \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
  1532
  \texttt{quant}, which may be, for example,
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1533
  \texttt{ALLGOALS} (all subgoals),
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1534
  \texttt{TRYALL} (try all subgoals, failure is OK),
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1535
  \texttt{FIRSTGOAL} (try subgoals until it succeeds once), 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1536
  \texttt{(fn tacf => tacf 4)} (subgoal 4), etc
16
5045dec52d2b polished
Christian Urban <urbanc@in.tum.de>
parents: 4
diff changeset
  1537
  (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
  1538
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1539
  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
  1540
  \begin{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1541
  Method.add_method : 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1542
  (bstring * (src -> Proof.context -> method) * string) -> theory trf ; 
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1543
  ( *
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1544
  * )
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1545
  \end{verbatim}
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1546
  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
  1547
  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
  1548
  (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
  1549
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1550
  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
  1551
  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
  1552
  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
  1553
  \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
  1554
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1555
2a69b119cdee added verbatim the notes by Jeremy
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1556
*}
75
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
  1557
(*>*)
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
  1558
end