CookBook/Readme.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 08 Jan 2009 22:47:15 +0000
changeset 64 9a6e5e0c4906
parent 60 5b9c6010897b
child 65 c8e9a4f97916
permissions -rw-r--r--
deleted old files and added code to give a special tag to the command ML
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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:
diff changeset
     1
theory Readme
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports Base
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
     5
chapter {* Comments for Authors *}
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:
diff changeset
     6
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
text {*
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
  \begin{itemize}
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
    10
  \item The Cookbook can be compiled on the command-line with:
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
    11
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
    12
  \begin{center}
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
    13
  @{text "isabelle make"}
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
    14
  \end{center}
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
    15
59
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
    16
  You very likely need a recent snapshot of Isabelle in order to compile
64
9a6e5e0c4906 deleted old files and added code to give a special tag to the command ML
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
    17
  the Cookbook. Some parts of the Cookbook also rely on compilation with
9a6e5e0c4906 deleted old files and added code to give a special tag to the command ML
Christian Urban <urbanc@in.tum.de>
parents: 60
diff changeset
    18
  PolyML.
59
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
    19
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    20
  \item You can include references to other Isabelle manuals using the 
47
4daf913fdbe1 hakked latex so that it does not display ML {* *}; general tuning
Christian Urban <urbanc@in.tum.de>
parents: 44
diff changeset
    21
  reference names from those manuals. To do this the following
54
1783211b3494 tuned; added document antiquotation ML_response_fake_both
Christian Urban <urbanc@in.tum.de>
parents: 53
diff changeset
    22
  four \LaTeX{} commands are defined:
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:
diff changeset
    23
  
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
  \begin{center}
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
  \begin{tabular}{l|c|c}
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
   & Chapters & Sections\\\hline
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    27
  Implementation Manual & @{text "\\ichcite{\<dots>}"} & @{text "\\isccite{\<dots>}"}\\
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    28
  Isar Reference Manual & @{text "\\rchcite{\<dots>}"} & @{text "\\rsccite{\<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:
diff changeset
    29
  \end{tabular}
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
  \end{center}
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    32
  So @{text "\\ichcite{ch:logic}"} yields a reference for the chapter about logic 
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:
diff changeset
    33
  in the implementation manual, namely \ichcite{ch:logic}.
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
  \item There are various document antiquotations defined for the 
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
    36
  Cookbook. They allow to check the written text against the current
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
    37
  Isabelle code and also allow to show responses of the ML-compiler.
52
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
    38
  Therefore authors are strongly encouraged to use antiquotations wherever
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
    39
  appropriate.
52
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
    40
  
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
    41
  The following antiquotations are defined:
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:
diff changeset
    42
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
  \begin{itemize}
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    44
  \item[$\bullet$] @{text "@{ML \"expr\" for vars in structs}"} should be used
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    45
  for displaying any ML-ex\-pression, because the antiquotation checks whether
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    46
  the expression is valid ML-code. The @{text "for"}- and @{text
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    47
  "in"}-arguments are optional. The former is used for evaluating open
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    48
  expressions by giving a list of free variables. The latter is used to
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    49
  indicate in which structure or structures the ML-expression should be
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    50
  evaluated. Examples are:
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
    51
  
54
1783211b3494 tuned; added document antiquotation ML_response_fake_both
Christian Urban <urbanc@in.tum.de>
parents: 53
diff changeset
    52
  \begin{center}\small
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    53
  \begin{tabular}{lll}
59
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
    54
  @{text "@{ML \"1 + 3\"}"}         &                 & @{ML "1 + 3"}\\
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
    55
  @{text "@{ML \"a + b\" for a b}"} & \;\;produce\;\; & @{ML "a + b" for a b}\\
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
    56
  @{text "@{ML Ident in OuterLex}"} &                 & @{ML Ident in OuterLex}\\
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    57
  \end{tabular}
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    58
  \end{center}
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    59
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    60
  \item[$\bullet$] @{text "@{ML_response \"expr\" \"pat\"}"} should be used to
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    61
  display ML-expressions and their response.  The first expression is checked
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    62
  like in the antiquotation @{text "@{ML \"expr\"}"}; the second is a pattern
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    63
  that specifies the result the first expression produces. This pattern can
59
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
    64
  contain @{text [quotes] "\<dots>"} for parts that you like to omit. The response of the
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    65
  first expression will be checked against this pattern. Examples are:
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    66
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    67
  \begin{center}\small
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
    68
  \begin{tabular}{l}
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    69
  @{text "@{ML_response \"1+2\" \"3\"}"}\\
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    70
  @{text "@{ML_response \"(1+2,3)\" \"(3,\<dots>)\"}"}
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
    71
  \end{tabular}
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
    72
  \end{center}
52
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
    73
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    74
  which produce respectively
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    75
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    76
  \begin{center}\small
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    77
  \begin{tabular}{p{3cm}p{3cm}}
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    78
  @{ML_response "1+2" "3"} &  
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    79
  @{ML_response "(1+2,3)" "(3,\<dots>)"}
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    80
  \end{tabular}
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    81
  \end{center}
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    82
  
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    83
  Note that this antiquotation can only be used when the result can be
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
    84
  constructed: it does not work when the code produces an exception or returns 
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
    85
  an abstract datatype (like @{ML_type thm} or @{ML_type cterm}).
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
    86
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    87
  \item[$\bullet$] @{text "@{ML_response_fake \"expr\" \"pat\"}"} works just
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    88
  like the antiquotation @{text "@{ML_response \"expr\" \"pat\"}"} above,
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    89
  except that the result-specification is not checked. Use this antiquotation
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    90
  when the result cannot be constructed or the code generates an
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    91
  exception. Examples are:
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    92
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    93
  \begin{center}\small
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    94
  \begin{tabular}{ll}
59
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
    95
  @{text "@{ML_response_fake"} & @{text "\"cterm_of @{theory} @{term \\\"a + b = c\\\"}\"}"}\\
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
    96
                               & @{text "\"a + b = c\"}"}\smallskip\\ 
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
    97
  @{text "@{ML_response_fake"} & @{text "\"($$ \\\"x\\\") (explode \\\"world\\\")\""}\\ 
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
    98
                               & @{text "\"Exception FAIL raised\"}"}
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
    99
  \end{tabular}
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   100
  \end{center}
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   101
59
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   102
  which produce respectively
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   103
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   104
  \begin{center}\small
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   105
  \begin{tabular}{p{7.2cm}}
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   106
  @{ML_response_fake "cterm_of @{theory} @{term \"a + b = c\"}" "a + b = c"}\smallskip\\
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   107
  @{ML_response_fake "($$ \"x\") (explode \"world\")" "Exception FAIL raised"}
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   108
  \end{tabular}
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   109
  \end{center}
54
1783211b3494 tuned; added document antiquotation ML_response_fake_both
Christian Urban <urbanc@in.tum.de>
parents: 53
diff changeset
   110
  
59
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   111
  This output mimics to some extend what the user sees when running the
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   112
  code.
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   113
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   114
  \item[$\bullet$] @{text "@{ML_response_fake_both \"expr\" \"pat\"}"} can be
54
1783211b3494 tuned; added document antiquotation ML_response_fake_both
Christian Urban <urbanc@in.tum.de>
parents: 53
diff changeset
   115
  used to show erroneous code. Neither the code nor the response will be
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   116
  checked. An example is:
52
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   117
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   118
  \begin{center}\small
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   119
  \begin{tabular}{ll}
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   120
  @{text "@{ML_response_fake_both"} & @{text "\"@{cterm \\\"1 + True\\\"}\""}\\
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   121
                                    & @{text "\"Type unification failed \<dots>\"}"}
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   122
  \end{tabular}
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   123
  \end{center}
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   124
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   125
  \item[$\bullet$] @{text "@{ML_file \"name\"}"} should be used when
59
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   126
  referring to a file. It checks whether the file exists.  An example
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   127
  is 
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   128
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   129
  @{text [display] "@{ML_file \"Pure/General/basics.ML\"}"}
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:
diff changeset
   130
  \end{itemize}
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   132
  The listed antiquotations honour options including @{text "[display]"} and 
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   133
  @{text "[quotes]"}. For example
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   134
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   135
  \begin{center}\small
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   136
  @{text "@{ML [quotes] \"\\\"foo\\\" ^ \\\"bar\\\"\"}"} \;\;produces\;\; @{text [quotes] "foobar"}
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   137
  \end{center}
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   138
59
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
   139
  whereas
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   140
  
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   141
  \begin{center}\small
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   142
  @{text "@{ML \"\\\"foo\\\" ^ \\\"bar\\\"\"}"} \;\;produces only\;\; @{text "foobar"}
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   143
  \end{center}
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   144
  
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   145
52
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   146
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
   147
  \item Functions and value bindings cannot be defined inside antiquotations; they need
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
   148
  to be included inside \isacommand{ML} \isa{\isacharverbatimopen \ldots \isacharverbatimclose}
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   149
  environments. In this way they are also checked by the compiler. Some \LaTeX-hack in 
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   150
  the Cookbook, however, ensures that the environment markers are not printed.
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:
diff changeset
   151
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   152
  \item Line numbers can be printed using 
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   153
  \isacommand{ML} \isa{\%linenumbers} \isa{\isacharverbatimopen \ldots \isacharverbatimclose} 
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   154
  for ML-code or \isacommand{lemma} \isa{\%linenumbers} @{text "..."} for proofs.
52
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
   155
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:
diff changeset
   156
  \end{itemize}
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
*}
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
dee4b3e66dfe added a readme chapter for prospective authors; added commands for referring to the Isar Reference Manual
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 47
diff changeset
   161
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:
diff changeset
   162
end