# HG changeset patch # User wenzelm # Date 933339572 -7200 # Node ID 320b412e5800687b539f8d08a5b660c2a6477757 # Parent 64c9f2364dae267a32a7bc33c3817d049de5afe6 more stuff; diff -r 64c9f2364dae -r 320b412e5800 doc-src/IsarRef/Makefile --- a/doc-src/IsarRef/Makefile Fri Jul 30 13:44:29 1999 +0200 +++ b/doc-src/IsarRef/Makefile Fri Jul 30 14:59:32 1999 +0200 @@ -14,8 +14,8 @@ NAME = isar-ref FILES = isar-ref.tex intro.tex basics.tex syntax.tex pure.tex \ - simplifier.tex classical.tex hol.tex ../isar.sty \ - ../rail.sty ../proof.sty ../iman.sty ../extra.sty ../manual.bib + clasimp.tex hol.tex ../isar.sty \ + ../rail.sty ../railsetup.sty ../proof.sty ../iman.sty ../extra.sty ../manual.bib dvi: $(NAME).dvi diff -r 64c9f2364dae -r 320b412e5800 doc-src/IsarRef/clasimp.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/clasimp.tex Fri Jul 30 14:59:32 1999 +0200 @@ -0,0 +1,7 @@ + +\chapter{The Simplifier and Classical Reasoner} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "isar-ref" +%%% End: diff -r 64c9f2364dae -r 320b412e5800 doc-src/IsarRef/classical.tex --- a/doc-src/IsarRef/classical.tex Fri Jul 30 13:44:29 1999 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,7 +0,0 @@ - -\chapter{The Classical Reasoner} - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "isar-ref" -%%% End: diff -r 64c9f2364dae -r 320b412e5800 doc-src/IsarRef/hol.tex --- a/doc-src/IsarRef/hol.tex Fri Jul 30 13:44:29 1999 +0200 +++ b/doc-src/IsarRef/hol.tex Fri Jul 30 14:59:32 1999 +0200 @@ -1,5 +1,5 @@ -\chapter{HOL specific elements} +\chapter{Isabelle/HOL specific elements} %%% Local Variables: %%% mode: latex diff -r 64c9f2364dae -r 320b412e5800 doc-src/IsarRef/isar-ref.tex --- a/doc-src/IsarRef/isar-ref.tex Fri Jul 30 13:44:29 1999 +0200 +++ b/doc-src/IsarRef/isar-ref.tex Fri Jul 30 14:59:32 1999 +0200 @@ -1,14 +1,25 @@ %% $Id$ -\documentclass[12pt]{report} -\usepackage{graphicx,a4,../iman,../extra,../proof,../rail,../isar,../pdfsetup} +\documentclass[12pt,fleqn]{report} +\usepackage{graphicx,a4,../iman,../extra,../proof,../rail,../railsetup,../isar,../pdfsetup} \title{\includegraphics[scale=0.5]{isabelle_isar} \\[4ex] The Isabelle/Isar Reference Manual} \author{\emph{Markus Wenzel} \\ TU M\"unchen} \makeindex +\railterm{lbrace,rbrace,llbrace,rrbrace} +\railterm{ident,longident,symident,var,textvar,typefree,typevar,nat,string,verbatim} + +\railalias{name}{\railqtoken{name}} +\railalias{nameref}{\railqtoken{nameref}} +\railalias{text}{\railqtoken{text}} +\railalias{type}{\railqtoken{type}} +\railalias{term}{\railqtoken{term}} +\railalias{prop}{\railqtoken{prop}} +\railalias{atom}{\railqtoken{atom}} + \setcounter{secnumdepth}{2} \setcounter{tocdepth}{2} @@ -16,20 +27,8 @@ \sloppy \binperiod %%%treat . like a binary operator -\railalias{lbrace}{\ttlbrace} -\railalias{rbrace}{\ttrbrace} -\railterm{lbrace,rbrace} - -\railterm{ident,longident,symident,var,textvar,typefree,typevar,nat,string,verbatim} -\railterm{name,nameref,text,type,term,prop,atom} +\renewcommand{\phi}{\varphi} -\makeatletter -\newcommand{\railtoken}[1]{{\rail@termfont{#1}}} -\newcommand{\railnonterm}[1]{{\rail@nontfont{#1}}} -\makeatother - -\newcommand\indexoutertoken[1]{\index{#1@\railtoken{#1} (outer syntax)|bold}} -\newcommand\indexouternonterm[1]{\index{#1@\railnonterm{#1} (outer syntax)|bold}} \begin{document} @@ -57,8 +56,7 @@ \include{basics} \include{syntax} \include{pure} -\include{simplifier} -\include{classical} +\include{clasimp} \include{hol} \begingroup diff -r 64c9f2364dae -r 320b412e5800 doc-src/IsarRef/pure.tex --- a/doc-src/IsarRef/pure.tex Fri Jul 30 13:44:29 1999 +0200 +++ b/doc-src/IsarRef/pure.tex Fri Jul 30 14:59:32 1999 +0200 @@ -1,6 +1,620 @@ \chapter{Common Isar elements} +FIXME $*$ indicates \emph{improper commands} + +\section{Theory commands} + +\subsection{Defining theories} + +\indexisarcmd{theory}\indexisarcmd{end}\indexisarcmd{context} +\begin{matharray}{rcl} + \isarcmd{theory} & : & \isartrans{\cdot}{theory} \\ + \isarcmd{context}^* & : & \isartrans{\cdot}{theory} \\ + \isarcmd{end} & : & \isartrans{theory}{\cdot} \\ +\end{matharray} + +Isabelle/Isar ``new-style'' theories are either defined via theory files or +interactively.\footnote{In contrast, ``old-style'' Isabelle theories support + batch processing only, with only the ML proof script part suitable for + interaction.} + +The first command of any theory has to be $\THEORY$, starting a new theory +based on the merge of existing ones. In interactive experiments, the theory +context may be changed by $\CONTEXT$ without creating a new theory. In both +cases the concluding command is $\END$, which has to be the very last one of +any proper theory file. + +\begin{rail} + 'theory' name '=' (name + '+') filespecs? ':' + ; + 'context' name + ; + 'end' + ;; + + filespecs : 'files' ((name | '(' name ')') +); +\end{rail} + +\begin{description} +\item [$\THEORY~A = B@1 + \cdots + B@n$] commences a new theory $A$ based on + existing ones $B@1 + \cdots + B@n$. Note that Isabelle's theory loader + system ensures that any of the base theories are properly loaded (and fully + up-to-date when $\THEORY$ is executed interactively). + +\item [$\CONTEXT~B$] enters existing theory context $B$, basically in + read-only mode, so only a limited set of commands may be performed. Just as + for $\THEORY$, the theory loader ensures that $B$ is loaded and up-to-date. + +\item [$\END$] concludes the current theory definition of context switch. +\end{description} + + +\subsection{Formal comments} + +\indexisarcmd{title}\indexisarcmd{chapter}\indexisarcmd{subsection}\indexisarcmd{subsubsection} +\indexisarcmd{text}\indexisarcmd{txt} +\begin{matharray}{rcl} + \isarcmd{title} & : & \isartrans{theory}{theory} \\ + \isarcmd{chapter} & : & \isartrans{theory}{theory} \\ + \isarcmd{subsection} & : & \isartrans{theory}{theory} \\ + \isarcmd{subsubsection} & : & \isartrans{theory}{theory} \\ + \isarcmd{text} & : & \isartrans{theory}{theory} \\ + \isarcmd{txt} & : & \isartrans{proof(state)}{proof(state)} \\ +\end{matharray} + +There are several commands to include \emph{formal comments} in theory and +proof specification. In contrast to source-level comments +\verb|(*|\dots\verb|*)|, which are stripped at the lexical level, any text +given as formal comment is meant to be part of the actual document. +Consequently, it would be included in the final printed version. + +Apart from plain prose, formal comments may also refer to logical entities of +the current theory context (types, terms, theorems etc.). Proper processing +of the text would then include some further consistency checks with the items +declared in the current theory, e.g.\ type-checking of included terms. +\footnote{The current version of Isabelle/Isar does not process formal + comments in any such way. This will be available as part of the automatic + theory and proof document preparation system (via (PDF)LaTeX) that is + planned for the near future.} + +\begin{rail} + 'title' text text? text? + ; + 'chapter' text + ; + 'subsection' text + ; + 'subsubsection' text + ; + 'text' text + ; + 'txt' text + ; +\end{rail} + +\begin{description} +\item [$\isarkeyword{title}~title~author~date$] specifies the document title + just as in typical LaTeX documents. +\item [$\isarkeyword{chapter}~text$, $\isarkeyword{subsection}~text$, + $\isarkeyword{subsubsection}~text$] specify chapter and subsection headings. +\item [$\TEXT~text$] specifies an actual body of prose text, including + references to formal entities.\footnote{The latter feature is not yet + exploited in any way.} +\item [$\TXT~text$] is similar to $\TEXT$, but may appear within proofs. +\end{description} + + +\subsection{Type classes and sorts} + +\indexisarcmd{classes}\indexisarcmd{classrel}\indexisarcmd{defaultsort} +\begin{matharray}{rcl} + \isarcmd{classes} & : & \isartrans{theory}{theory} \\ + \isarcmd{classrel} & : & \isartrans{theory}{theory} \\ + \isarcmd{defaultsort} & : & \isartrans{theory}{theory} \\ +\end{matharray} + +\begin{rail} + 'classes' (name ('<' (nameref ',' +))? comment? +) + ; + 'classrel' nameref '<' nameref comment? + ; + 'defaultsort' sort comment? + ; +\end{rail} + +\begin{description} +\item [$\isarkeyword{classes}~c' | '<=') transpat comment? +) + ; + transpat: ('(' nameref ')')? string + ; +\end{rail} + +\begin{description} +\item [$\isarkeyword{syntax}~mode~decls$] is similar to $\CONSTS~decls$, + except the actual logical signature extension. Thus the context free + grammar of Isabelle's inner syntax may be augmented in arbitrary ways. The + $mode$ argument refers to the print mode that the grammar rules belong; + unless there is the \texttt{output} flag given, all productions are added + both to the input and output grammar. +\item [$\isarkeyword{translations}~rule~\dots$] specifies syntactic + translation rules (macros): parse/print rules (\texttt{==}), parse rules + (\texttt{=>}), print rules (\texttt{<=}). Translation patterns may be + prefixed by the syntactic category to be used for parsing; the default is + \texttt{logic}. +\end{description} + + +\subsection{Axioms and theorems} + +\indexisarcmd{axioms}\indexisarcmd{theorems}\indexisarcmd{lemmas} +\begin{matharray}{rcl} + \isarcmd{axioms} & : & \isartrans{theory}{theory} \\ + \isarcmd{theorems} & : & \isartrans{theory}{theory} \\ + \isarcmd{lemmas} & : & \isartrans{theory}{theory} \\ +\end{matharray} + +\begin{rail} + 'axioms' (name attributes? ':' prop comment? +) + ; + ('theorems' | 'lemmas') thmdef? thmrefs + ; +\end{rail} + +\begin{description} +\item [$\isarkeyword{axioms}~name: \phi~\dots$] introduces arbitrary + statements as logical axioms. In fact, axioms are ``axiomatic theorems'', + and may be referred as any other theorems later. + + Axioms are usually only introduced when declaring new logical systems. + Everyday work is normally done the hard way, with proper definitions and + actual theorems. +\item [$\isarkeyword{theorems}~name = thms$] stores lists of existing theorems + as $name$. Typical applications would also involve attributes to augment + the default simpset, for example. +\item [$\isarkeyword{lemmas}$] is similar to $\isarkeyword{theorems}$, but + tags the results as ``lemma''. +\end{description} + + +\subsection{Manipulating name spaces} + +\indexisarcmd{global}\indexisarcmd{local}\indexisarcmd{path} +\begin{matharray}{rcl} + \isarcmd{global} & : & \isartrans{theory}{theory} \\ + \isarcmd{local} & : & \isartrans{theory}{theory} \\ + \isarcmd{path} & : & \isartrans{theory}{theory} \\ +\end{matharray} + +\begin{rail} + 'global' + ; + 'local' + ; + 'path' nameref + ; +\end{rail} + +\begin{description} +\item [$\isarkeyword{global}$] FIXME +\item [$\isarkeyword{local}$] FIXME +\item [$\isarkeyword{path}~name$] FIXME +\end{description} + + +\subsection{Incorporating ML code} + +\indexisarcmd{use}\indexisarcmd{ML}\indexisarcmd{setup} +\begin{matharray}{rcl} + \isarcmd{use} & : & \isartrans{\cdot}{\cdot} \\ + \isarcmd{ML} & : & \isartrans{\cdot}{\cdot} \\ + \isarcmd{setup} & : & \isartrans{\cdot}{\cdot} \\ +\end{matharray} + +\begin{rail} + 'use' name + ; + 'ML' text + ; + 'setup' text + ; +\end{rail} + +\begin{description} +\item [$\isarkeyword{use}~file$] FIXME +\item [$\isarkeyword{ML}~text$] FIXME +\item [$\isarkeyword{setup}~text$] FIXME +\end{description} + + +\subsection{ML translation functions} + +\indexisarcmd{parse_ast_translation}\indexisarcmd{parse_translation} +\indexisarcmd{print_translation}\indexisarcmd{typed_print_translation} +\indexisarcmd{print_ast_translation}\indexisarcmd{token_translation} +\begin{matharray}{rcl} + \isarcmd{parse_ast_translation} & : & \isartrans{theory}{theory} \\ + \isarcmd{parse_translation} & : & \isartrans{theory}{theory} \\ + \isarcmd{print_translation} & : & \isartrans{theory}{theory} \\ + \isarcmd{typed_print_translation} & : & \isartrans{theory}{theory} \\ + \isarcmd{print_ast_translation} & : & \isartrans{theory}{theory} \\ + \isarcmd{token_translation} & : & \isartrans{theory}{theory} \\ +\end{matharray} + +Syntax translation functions written in ML admit almost arbitrary +manipulations of Isabelle's inner syntax. Any of the above commands have a +single \railqtoken{text} argument that refers to an ML expression of +appropriate type. See \cite[\S8]{isabelle-ref} for more information on syntax +transformations. + + +\subsection{Oracles} + +\indexisarcmd{oracle} +\begin{matharray}{rcl} + \isarcmd{oracle} & : & \isartrans{theory}{theory} \\ +\end{matharray} + +\begin{rail} + 'oracle' name '=' text comment? + ; +\end{rail} + +\begin{description} +\item [$\isarkeyword{oracle}~name=text$] FIXME +\end{description} + + +\section{Proof commands} + +\subsection{Goal statements} + +\indexisarcmd{} +\begin{matharray}{rcl} + \isarcmd{theorem} & : & \isartrans{theory}{proof(prove)} \\ + \isarcmd{lemma} & : & \isartrans{theory}{proof(prove)} \\ + \isarcmd{have} & : & \isartrans{proof(state)}{proof(prove)} \\ + \isarcmd{show} & : & \isartrans{proof(state)}{proof(prove)} \\ + \isarcmd{hence} & : & \isartrans{proof(state)}{proof(prove)} \\ + \isarcmd{thus} & : & \isartrans{proof(state)}{proof(prove)} \\ +\end{matharray} + +\begin{rail} + ('theorem' | 'lemma') goal + ; + ('have' | 'show' | 'hence' | 'thus') goal + ; + + goal: thmdecl? proppat comment? + ; +\end{rail} + +\begin{description} +\item [$\THEOREM{name}{\phi}$] enters proof mode with $\phi$ as main goal, + eventually resulting in some theorem $\turn \phi$, which stored in the + theory. +\item [$\LEMMANAME$] is similar to $\THEOREMNAME$, but tags the result as + ``lemma''. +\item [$\HAVE{name}{\phi}$] FIXME +\item [$\SHOW{name}{\phi}$] FIXME +\item [$\HENCE{name}{\phi}$] FIXME +\item [$\THUS{name}{\phi}$] FIXME +\end{description} + + +\subsection{Initial and terminal proof steps} + +\indexisarcmd{proof}\indexisarcmd{qed}\indexisarcmd{by} +\indexisarcmd{.}\indexisarcmd{..}\indexisarcmd{sorry} +\begin{matharray}{rcl} + \isarcmd{proof} & : & \isartrans{proof(prove)}{proof(state)} \\ + \isarcmd{qed} & : & \isartrans{proof(state)}{proof(state) ~|~ theory} \\ + \isarcmd{by} & : & \isartrans{proof(prove)}{proof(state) ~|~ theory} \\ + \isarcmd{.} & : & \isartrans{proof(prove)}{proof(state) ~|~ theory} \\ + \isarcmd{..} & : & \isartrans{proof(prove)}{proof(state) ~|~ theory} \\ + \isarcmd{sorry} & : & \isartrans{proof(prove)}{proof(state) ~|~ theory} \\ +\end{matharray} + +\begin{rail} + 'proof' interest? meth? comment? + ; + 'qed' meth? comment? + ; + 'by' meth meth? comment? + ; + ('.' | '..' | 'sorry') comment? + ; + + meth: method interest? + ; +\end{rail} + +\begin{description} +\item [$ $] FIXME +\end{description} + + +\subsection{Facts and forward chaining} + +\indexisarcmd{note}\indexisarcmd{then}\indexisarcmd{from}\indexisarcmd{with} +\begin{matharray}{rcl} + \isarcmd{note} & : & \isartrans{proof(state)}{proof(state)} \\ + \isarcmd{then} & : & \isartrans{proof(state)}{proof(chain)} \\ + \isarcmd{from} & : & \isartrans{proof(state)}{proof(chain)} \\ + \isarcmd{with} & : & \isartrans{proof(state)}{proof(chain)} \\ +\end{matharray} + +\begin{rail} + 'note' thmdef? thmrefs comment? + ; + 'then' comment? + ; + ('from' | 'with') thmrefs comment? + ; +\end{rail} + +\begin{description} +\item [$ $] FIXME +\end{description} + + +\subsection{Proof context} + +\indexisarcmd{fix}\indexisarcmd{assume}\indexisarcmd{presume}\indexisarcmd{def}\indexisarcmd{let} +\begin{matharray}{rcl} + \isarcmd{fix} & : & \isartrans{proof(state)}{proof(state)} \\ + \isarcmd{assume} & : & \isartrans{proof(state)}{proof(state)} \\ + \isarcmd{presume} & : & \isartrans{proof(state)}{proof(state)} \\ + \isarcmd{def} & : & \isartrans{proof(state)}{proof(state)} \\ + \isarcmd{let} & : & \isartrans{proof(state)}{proof(state)} \\ +\end{matharray} + +\begin{rail} + 'fix' (var +) comment? + ; + ('assume' | 'presume') thmdecl? (proppat +) comment? + ; + 'def' thmdecl? var '==' termpat comment? + ; + 'let' ((term + 'as') '=' term comment? + 'and') + ; + + var: name ('::' type)? + ; +\end{rail} + +\begin{description} +\item [$ $] FIXME +\end{description} + + +\subsection{Block structure} + +\indexisarcmd{next}\indexisarcmd{\{\{}\indexisarcmd{\}\}} +\begin{matharray}{rcl} + \isarcmd{next} & : & \isartrans{proof(state)}{proof(state)} \\ + \isarcmd{\{\{} & : & \isartrans{proof(state)}{proof(state)} \\ + \isarcmd{\}\}} & : & \isartrans{proof(state)}{proof(state)} \\ +\end{matharray} + +\begin{rail} + llbrace + ; + rrbrace + ; + 'next' + ; +\end{rail} + +\begin{description} +\item [$ $] FIXME +\end{description} + + +\subsection{Calculational proof} + +\indexisarcmd{also}\indexisarcmd{finally} +\begin{matharray}{rcl} + \isarcmd{also} & : & \isartrans{proof(state)}{proof(state)} \\ + \isarcmd{finally} & : & \isartrans{proof(state)}{proof(chain)} \\ +\end{matharray} + +\begin{rail} + ('also' | 'finally') transrules? comment? + ; + + transrules: '(' thmrefs ')' interest? + ; +\end{rail} + +\begin{description} +\item [$ $] FIXME +\end{description} + + + +\subsection{Improper proof steps} + +\indexisarcmd{apply}\indexisarcmd{then_apply}\indexisarcmd{back} +\begin{matharray}{rcl} + \isarcmd{apply}^* & : & \isartrans{proof}{proof} \\ + \isarcmd{then_apply}^* & : & \isartrans{proof}{proof} \\ + \isarcmd{back}^* & : & \isartrans{proof}{proof} \\ +\end{matharray} + +\railalias{thenapply}{then\_apply} +\railterm{thenapply} + +\begin{rail} + 'apply' method + ; + thenapply method + ; + 'back' + ; +\end{rail} + +\begin{description} +\item [$ $] FIXME +\end{description} + + +\section{Other commands} + +\subsection{Diagnostics} + +\indexisarcmd{typ}\indexisarcmd{term}\indexisarcmd{prop}\indexisarcmd{thm} +\begin{matharray}{rcl} + \isarcmd{typ} & : & \isarkeep{theory~|~proof} \\ + \isarcmd{term} & : & \isarkeep{theory~|~proof} \\ + \isarcmd{prop} & : & \isarkeep{theory~|~proof} \\ + \isarcmd{thm} & : & \isarkeep{theory~|~proof} \\ +\end{matharray} + +\begin{rail} + 'typ' type + ; + 'term' term + ; + 'prop' prop + ; + 'thm' thmrefs + ; +\end{rail} + +\begin{description} +\item [$\isarkeyword{typ}~\tau$, $\isarkeyword{term}~t$, + $\isarkeyword{prop}~\phi$] read and print types / terms / propositions + according to the current theory or proof context. +\item [$\isarkeyword{thm}~thms$] retrieves lists of theorems from the current + theory or proof context. Note that any attributes included in the theorem + specifications are applied to a temporary proof context derived from the + current theory or proof; the resulting context is discarded. +\end{description} + + +\subsection{System operations} + +\indexisarcmd{cd}\indexisarcmd{pwd}\indexisarcmd{use_thy}\indexisarcmd{use_thy_only} +\indexisarcmd{update_thy}\indexisarcmd{update_thy_only} +\begin{matharray}{rcl} + \isarcmd{cd} & : & \isarkeep{\cdot} \\ + \isarcmd{pwd} & : & \isarkeep{\cdot} \\ + \isarcmd{use_thy} & : & \isarkeep{\cdot} \\ + \isarcmd{use_thy_only} & : & \isarkeep{\cdot} \\ + \isarcmd{update_thy} & : & \isarkeep{\cdot} \\ + \isarcmd{update_thy_only} & : & \isarkeep{\cdot} \\ +\end{matharray} + +\begin{description} +\item [$\isarkeyword{cd}~name$] changes the current directory of the Isabelle + process. +\item [$\isarkeyword{pwd}~$] prints the current working directory. +\item [$\isarkeyword{use_thy}~name$, $\isarkeyword{use_thy_only}~name$, + $\isarkeyword{update_thy}~name$, $\isarkeyword{update_thy_only}~name$] load + theory files. These commands are exactly the same as the corresponding ML + functions (see also \cite[\S1 and \S6]{isabelle-ref}). +\end{description} + + %%% Local Variables: %%% mode: latex %%% TeX-master: "isar-ref" diff -r 64c9f2364dae -r 320b412e5800 doc-src/IsarRef/simplifier.tex --- a/doc-src/IsarRef/simplifier.tex Fri Jul 30 13:44:29 1999 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,7 +0,0 @@ - -\chapter{The Simplifier} - -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "isar-ref" -%%% End: diff -r 64c9f2364dae -r 320b412e5800 doc-src/IsarRef/syntax.tex --- a/doc-src/IsarRef/syntax.tex Fri Jul 30 13:44:29 1999 +0200 +++ b/doc-src/IsarRef/syntax.tex Fri Jul 30 14:59:32 1999 +0200 @@ -5,6 +5,8 @@ \chapter{Isar document syntax} +FIXME shortcut + FIXME important note: inner versus outer syntax \section{Lexical matters} @@ -12,22 +14,26 @@ \section{Common syntax entities} The Isar proof and theory language syntax has been carefully designed with -orthogonality in mind. Many common syntax entities such that those for names, -terms, types etc.\ have been factored out. Some of these basic syntactic -entities usually represent the level of abstraction for error messages: e.g.\ -some higher syntax element such as $\CONSTS$ referring to \railtoken{name} or -\railtoken{type}, would really report a missing \railtoken{name} or -\railtoken{type} rather than any of its constituent primitive tokens (as -defined below). These quasi-tokens are represented in the syntax diagrams -below using the same font as actual tokens (such as \railtoken{string}). +orthogonality in mind. Subsequently, we introduce several basic syntactic +entities, such as names, terms, theorem specifications, which have been +factored out of the actual Isar language elements described later. + +Some of the basic syntactic entities introduced below act much like tokens +rather than nonterminals, in particular for error messages are concerned. +E.g.\ syntax elements such as $\CONSTS$ referring to \railqtoken{name} or +\railqtoken{type} would really report a missing \railqtoken{name} or +\railqtoken{type} rather than any of its constituent primitive tokens +(\railtoken{ident}, \railtoken{string} etc.). \subsection{Names} -Entity \railtoken{name} usually refers to any name of types, constants, +Entity \railqtoken{name} usually refers to any name of types, constants, theorems, etc.\ to be \emph{declared} or \emph{defined} (so qualified -identifiers are excluded). Already existing objects are typically referenced -by \railtoken{nameref}. +identifiers are excluded). Quoted strings provide an escape for +non-identifier names or those ruled out by outer syntax keywords (e.g.\ +\verb|"let"|). Already existing objects are usually referenced by +\railqtoken{nameref}. \indexoutertoken{name}\indexoutertoken{nameref} \begin{rail} @@ -40,23 +46,23 @@ \subsection{Comments} -Large chunks of verbatim \railtoken{text} are usually given -\railtoken{verbatim}, i.e.\ enclosed in \verb|{*|~\verb|*}|; for convenience, -any of the smaller text entities (\railtoken{ident}, \railtoken{string} etc.) -are admitted as well. Almost any of the Isar commands may be annotated by a -marginal comment: \texttt{--} \railtoken{text}. Note that this kind of -comment is actually part of the language, while source level comments -\verb|(*|~\verb|*)| are already stripped at the lexical level. A few commands -such as $\PROOFNAME$ admit some parts to be mark with a ``level of interest'': -currently only \texttt{\%} for ``boring, don't read this''. +Large chunks of plain \railqtoken{text} are usually given \railtoken{verbatim}, +i.e.\ enclosed in \verb|{*|\dots\verb|*}|. For convenience, any of the +smaller text entities (\railtoken{ident}, \railtoken{string} etc.) are +admitted as well. Almost any of the Isar commands may be annotated by a +marginal \railnonterm{comment}: \texttt{--} \railqtoken{text}. Note that this +kind of comment is actually part of the language, while source level comments +\verb|(*|\dots\verb|*)| are already stripped at the lexical level. A few +commands such as $\PROOFNAME$ admit additional markup with a ``level of +interest'', currently only \texttt{\%} for ``boring, don't read this''. \indexoutertoken{text}\indexouternonterm{comment}\indexouternonterm{interest} \begin{rail} text : verbatim | nameref ; - comment : (() | '--' text) + comment : '--' text ; - interest : (() | '\%') + interest : '\%' ; \end{rail} @@ -84,16 +90,16 @@ flexible in order to be modeled explicitly at the outer theory level. Basically, any such entity would have to be quoted at the outer level to turn it into a single token, with the actual parsing deferred to some functions -that read and type-check terms etc.\ (note that \railtoken{prop}s will be -handled differently from plain \railtoken{term}s here). For convenience, the +that read and type-check terms etc.\ (note that \railqtoken{prop}s will be +handled differently from plain \railqtoken{term}s here). For convenience, the quotes may be omitted for any \emph{atomic} term or type (e.g.\ a single variable). \indexoutertoken{type}\indexoutertoken{term}\indexoutertoken{prop} \begin{rail} - type : ident | longident | symident | typefree | typevar | string + type : nameref | typefree | typevar ; - term : ident | longident | symident | var | textvar | nat | string + term : nameref | var | textvar | nat ; prop : term ; @@ -116,7 +122,7 @@ Statements like $\SHOWNAME$ involve propositions, some others like $\DEFNAME$ plain terms. Any of these usually admit automatic binding of schematic text variables by giving (optional) patterns $\IS{p@1 \dots p@n}$. For -\railtoken{prop}s the $\CONCLNAME$ part refers to the conclusion only, in case +\railqtoken{prop}s the $\CONCLNAME$ part refers to the conclusion only, in case actual rules are involved, rather than atomic propositions. \indexouternonterm{termpat}\indexouternonterm{proppat} @@ -130,9 +136,9 @@ \subsection{Mixfix annotations} -Mixfix annotations specify concrete \emph{inner} syntax. Some commands such -as $\TYPES$ admit infixes only, while $\CONSTS$ etc.\ support the full range -of general mixfixes and binders. +Mixfix annotations specify concrete \emph{inner} syntax of Isabelle types and +terms. Some commands such as $\TYPES$ admit infixes only, while $\CONSTS$ +etc.\ support the full range of general mixfixes and binders. \indexouternonterm{infix}\indexouternonterm{mixfix} \begin{rail} @@ -145,39 +151,45 @@ \end{rail} -\subsection{Attributes and theorem specifications}\label{sec:syn-att} +\subsection{Attributes and theorems}\label{sec:syn-att} Attributes (and proof methods, see \S\ref{sec:syn-meth}) have their own ``semi-inner'' syntax, which does not have to be atomic at the outer level unlike that of types and terms. Instead, the attribute argument specifications may be any sequence of atomic entities (identifiers, strings -etc.), or properly bracketed argument lists. Below \railtoken{atom} refers to +etc.), or properly bracketed argument lists. Below \railqtoken{atom} refers to any atomic entity (\railtoken{ident}, \railtoken{longident}, \railtoken{symident} etc.), including keywords that conform to \railtoken{symident}, but do not coincide with actual command names. \indexoutertoken{atom}\indexouternonterm{args}\indexouternonterm{attributes} \begin{rail} - args : (atom | '(' (args *) ')' | '[' (args *) ']' | lbrace (args *) rbrace) * + atom : nameref | typefree | typevar | var | textvar | nat + ; + arg : atom | '(' args ')' | '[' args ']' | lbrace args rbrace ; - attributes : '[' (name args + ',') ']' + args : arg * + ; + attributes : '[' (nameref args * ',') ']' ; \end{rail} -Theorem specifications come in three flavours: \railnonterm{thmdecl} usually -refers to the result of a goal statement (such as $\SHOWNAME$), +Theorem specifications come in three flavors: \railnonterm{thmdecl} usually +refers to the result of an assumption or goal statement (e.g.\ $\SHOWNAME$), \railnonterm{thmdef} collects lists of existing theorems (as in $\NOTENAME$), -\railnonterm{thmref} refers to any list of existing theorems (e.g.\ occurring +\railnonterm{thmrefs} refers to any list of existing theorems (e.g.\ occurring as proof method arguments). Any of these may include lists of attributes, which are applied to the preceding theorem or list of theorems. -\indexouternonterm{thmdecl}\indexouternonterm{thmdef}\indexouternonterm{thmref} +\indexouternonterm{thmdecl}\indexouternonterm{thmdef}\indexouternonterm{thmrefs} \begin{rail} - thmdecl : (() | name) (() | attributes) ':' + thmname : name attributes | name | attributes ; - thmdef : (() | name) (() | attributes) '=' + thmdecl : thmname ':' ; - thmref : (name (() | attributes) +) + thmdef : thmname '=' + ; + thmrefs : nameref (() | attributes) + ; \end{rail} @@ -187,15 +199,18 @@ Proof methods are either basic ones, or expressions composed of methods via ``\texttt{,}'' (sequential composition), ``\texttt{|}'' (alternatives), ``\texttt{?}'' (try), ``\texttt{*}'' (repeat, ${} \ge 0$ times), -``\texttt{+}'' (repeat, ${} > 0$ times). In practice, proof methods are -typically just a comma separeted list of \railtoken{name}~\railtoken{args} +``\texttt{+}'' (repeat, ${} > 0$ times). In practice, proof methods are very +often just a comma separated list of \railqtoken{nameref}~\railnonterm{args} specifications. Thus the syntax is similar to that of attributes, with plain parentheses instead of square brackets (see also \S\ref{sec:syn-att}). Note -that parentheses may be dropped for methods without arguments. +that parentheses may be dropped for single method specifications without +arguments. \indexouternonterm{method} \begin{rail} - method : (name args | (name | '(' method ')') (() | '?' | '*' | '+')) + (',' | '|') + method : (nameref | '(' methods ')') (() | '?' | '*' | '+') + ; + methods : (nameref args | method) + (',' | '|') ; \end{rail}