doc-src/Ref/theory-syntax.tex
author lcp
Tue, 12 Jul 1994 18:05:03 +0200
changeset 467 92868dab2939
parent 325 49baeba86546
child 1080 13c35eb5169b
permissions -rw-r--r--
new cardinal arithmetic developments

%% $Id$

\appendix
\newlinechar=-1  %mathsing.sty sets \newlinechar=`\|, which would cause mayhem

\chapter{Syntax of Isabelle Theories}\label{app:TheorySyntax}
Chapter~\ref{sec:ref-defining-theories} explains the meanings of these
constructs.  The syntax obeys the following conventions:
\begin{itemize}
\item {\tt Typewriter font} denotes terminal symbols.
\item $id$, $tid$, $nat$, $string$ and $text$ are the lexical classes of
  identifiers, type identifiers, natural numbers, \ML\ strings (with their
  quotation marks) and arbitrary \ML\ text.  The first three are fully defined
  in 
\iflabelundefined{Defining-Logics}%
    {{\it The Isabelle Reference Manual}, chapter `Defining Logics'}%
    {Chap.\ts\ref{Defining-Logics}}.
\end{itemize}
Comments in theories take the form {\tt (*} {\it text\/} {\tt*)}, where
{\it text\/} should not contain the character sequence {\tt*)}.

\begin{rail}

theoryDef : id '=' (name + '+') ('+' extension | ());

name: id | string;

extension : classes default types arities consts trans rules \\ 'end' ml
          ;

classes : ()
        | 'classes' ( ( id (  ()
                            | '<' (id + ',')
                           ) 
                       ) + )
        ;

default : ()
        | 'default' sort 
        ;

sort :  id
     | '\{' (id * ',') '\}'
     ;

types :  ()
      | 'types' ( ( typeDecl ( () | '(' infix ')' ) ) + )
      ;

typeDecl : ( () | tid | '(' ( tid + ',' ) ')') name ( () | '='  string );

infix : ( 'infixr' | 'infixl' ) nat;


arities :  ()
        | 'arities' ((( name + ',' ) '::' arity ) + )
        ;

arity   : ( () 
          | '(' (sort + ',') ')' 
          ) id
        ;


consts :  ()
       | 'consts' ( ( constDecl ( () | ( '(' mixfix ')' ) ) ) + )
       ;

constDecl : ( name + ',') '::' string ;


mixfix :  string ( () | ( () | ('[' (nat + ',') ']')) nat )
       | infix
       | 'binder' string nat ;

trans : ()
      | 'translations' ( pat ( '==' | '=>' | '<=' ) pat + )
      ;

pat : ( () | ( '(' id ')' ) ) string;

rules :  ()
      | 'rules' (( id string ) + )
      ;

ml :  ()
   | 'ML' text
   ;

\end{rail}