# HG changeset patch # User wenzelm # Date 1124484241 -7200 # Node ID 096792bdc58e80921df3c2502b678a637aa4a4b3 # Parent 153fe83804c93e8c56fe477646c482ecc6a9329b tuned generated stuff; diff -r 153fe83804c9 -r 096792bdc58e doc-src/AxClass/Group/document/Group.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/AxClass/Group/document/Group.tex Fri Aug 19 22:44:01 2005 +0200 @@ -0,0 +1,511 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Group}% +\isamarkuptrue% +% +\isamarkupheader{Basic group theory% +} +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isamarkupfalse% +\isacommand{theory}\ Group\ \isakeyword{imports}\ Main\ \isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% +% +\begin{isamarkuptext}% +\medskip\noindent The meta-level type system of Isabelle supports + \emph{intersections} and \emph{inclusions} of type classes. These + directly correspond to intersections and inclusions of type + predicates in a purely set theoretic sense. This is sufficient as a + means to describe simple hierarchies of structures. As an + illustration, we use the well-known example of semigroups, monoids, + general groups and Abelian groups.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Monoids and Groups% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +First we declare some polymorphic constants required later for the + signature parts of our structures.% +\end{isamarkuptext}% +\isamarkupfalse% +\isacommand{consts}\isanewline +\ \ times\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isasymodot}{\isachardoublequote}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline +\ \ invers\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\ \ \ \ {\isacharparenleft}{\isachardoublequote}{\isacharparenleft}{\isacharunderscore}{\isasyminv}{\isacharparenright}{\isachardoublequote}\ {\isacharbrackleft}{\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isadigit{0}}{\isacharbrackright}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharparenright}\isanewline +\ \ one\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ \ \ \ {\isacharparenleft}{\isachardoublequote}{\isasymone}{\isachardoublequote}{\isacharparenright}\isamarkuptrue% +% +\begin{isamarkuptext}% +\noindent Next we define class \isa{monoid} of monoids with + operations \isa{{\isasymodot}} and \isa{{\isasymone}}. Note that multiple class + axioms are allowed for user convenience --- they simply represent + the conjunction of their respective universal closures.% +\end{isamarkuptext}% +\isamarkupfalse% +\isacommand{axclass}\ monoid\ {\isasymsubseteq}\ type\isanewline +\ \ assoc{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x\ {\isasymodot}\ y{\isacharparenright}\ {\isasymodot}\ z\ {\isacharequal}\ x\ {\isasymodot}\ {\isacharparenleft}y\ {\isasymodot}\ z{\isacharparenright}{\isachardoublequote}\isanewline +\ \ left{\isacharunderscore}unit{\isacharcolon}\ {\isachardoublequote}{\isasymone}\ {\isasymodot}\ x\ {\isacharequal}\ x{\isachardoublequote}\isanewline +\ \ right{\isacharunderscore}unit{\isacharcolon}\ {\isachardoublequote}x\ {\isasymodot}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequote}\isamarkuptrue% +% +\begin{isamarkuptext}% +\noindent So class \isa{monoid} contains exactly those types + \isa{{\isasymtau}} where \isa{{\isasymodot}\ {\isasymColon}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}} and \isa{{\isasymone}\ {\isasymColon}\ {\isasymtau}} + are specified appropriately, such that \isa{{\isasymodot}} is associative and + \isa{{\isasymone}} is a left and right unit element for the \isa{{\isasymodot}} + operation.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\begin{isamarkuptext}% +\medskip Independently of \isa{monoid}, we now define a linear + hierarchy of semigroups, general groups and Abelian groups. Note + that the names of class axioms are automatically qualified with each + class name, so we may re-use common names such as \isa{assoc}.% +\end{isamarkuptext}% +\isamarkupfalse% +\isacommand{axclass}\ semigroup\ {\isasymsubseteq}\ type\isanewline +\ \ assoc{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x\ {\isasymodot}\ y{\isacharparenright}\ {\isasymodot}\ z\ {\isacharequal}\ x\ {\isasymodot}\ {\isacharparenleft}y\ {\isasymodot}\ z{\isacharparenright}{\isachardoublequote}\isanewline +\isanewline +\isamarkupfalse% +\isacommand{axclass}\ group\ {\isasymsubseteq}\ semigroup\isanewline +\ \ left{\isacharunderscore}unit{\isacharcolon}\ {\isachardoublequote}{\isasymone}\ {\isasymodot}\ x\ {\isacharequal}\ x{\isachardoublequote}\isanewline +\ \ left{\isacharunderscore}inverse{\isacharcolon}\ {\isachardoublequote}x{\isasyminv}\ {\isasymodot}\ x\ {\isacharequal}\ {\isasymone}{\isachardoublequote}\isanewline +\isanewline +\isamarkupfalse% +\isacommand{axclass}\ agroup\ {\isasymsubseteq}\ group\isanewline +\ \ commute{\isacharcolon}\ {\isachardoublequote}x\ {\isasymodot}\ y\ {\isacharequal}\ y\ {\isasymodot}\ x{\isachardoublequote}\isamarkuptrue% +% +\begin{isamarkuptext}% +\noindent Class \isa{group} inherits associativity of \isa{{\isasymodot}} + from \isa{semigroup} and adds two further group axioms. Similarly, + \isa{agroup} is defined as the subset of \isa{group} such that + for all of its elements \isa{{\isasymtau}}, the operation \isa{{\isasymodot}\ {\isasymColon}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}} is even commutative.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Abstract reasoning% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +In a sense, axiomatic type classes may be viewed as \emph{abstract + theories}. Above class definitions gives rise to abstract axioms + \isa{assoc}, \isa{left{\isacharunderscore}unit}, \isa{left{\isacharunderscore}inverse}, \isa{commute}, where any of these contain a type variable \isa{{\isacharprime}a\ {\isasymColon}\ c} that is restricted to types of the corresponding class \isa{c}. \emph{Sort constraints} like this express a logical + precondition for the whole formula. For example, \isa{assoc} + states that for all \isa{{\isasymtau}}, provided that \isa{{\isasymtau}\ {\isasymColon}\ semigroup}, the operation \isa{{\isasymodot}\ {\isasymColon}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}} is associative. + + \medskip From a technical point of view, abstract axioms are just + ordinary Isabelle theorems, which may be used in proofs without + special treatment. Such ``abstract proofs'' usually yield new + ``abstract theorems''. For example, we may now derive the following + well-known laws of general groups.% +\end{isamarkuptext}% +\isamarkupfalse% +\isacommand{theorem}\ group{\isacharunderscore}right{\isacharunderscore}inverse{\isacharcolon}\ {\isachardoublequote}x\ {\isasymodot}\ x{\isasyminv}\ {\isacharequal}\ {\isacharparenleft}{\isasymone}{\isasymColon}{\isacharprime}a{\isasymColon}group{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkupfalse% +\isacommand{proof}\ {\isacharminus}\isanewline +\ \ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}x\ {\isasymodot}\ x{\isasyminv}\ {\isacharequal}\ {\isasymone}\ {\isasymodot}\ {\isacharparenleft}x\ {\isasymodot}\ x{\isasyminv}{\isacharparenright}{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isachardot}left{\isacharunderscore}unit{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{also}\ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isasymone}\ {\isasymodot}\ x\ {\isasymodot}\ x{\isasyminv}{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ semigroup{\isachardot}assoc{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{also}\ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isacharparenleft}x{\isasyminv}{\isacharparenright}{\isasyminv}\ {\isasymodot}\ x{\isasyminv}\ {\isasymodot}\ x\ {\isasymodot}\ x{\isasyminv}{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isachardot}left{\isacharunderscore}inverse{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{also}\ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isacharparenleft}x{\isasyminv}{\isacharparenright}{\isasyminv}\ {\isasymodot}\ {\isacharparenleft}x{\isasyminv}\ {\isasymodot}\ x{\isacharparenright}\ {\isasymodot}\ x{\isasyminv}{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ semigroup{\isachardot}assoc{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{also}\ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isacharparenleft}x{\isasyminv}{\isacharparenright}{\isasyminv}\ {\isasymodot}\ {\isasymone}\ {\isasymodot}\ x{\isasyminv}{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isachardot}left{\isacharunderscore}inverse{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{also}\ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isacharparenleft}x{\isasyminv}{\isacharparenright}{\isasyminv}\ {\isasymodot}\ {\isacharparenleft}{\isasymone}\ {\isasymodot}\ x{\isasyminv}{\isacharparenright}{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ semigroup{\isachardot}assoc{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{also}\ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isacharparenleft}x{\isasyminv}{\isacharparenright}{\isasyminv}\ {\isasymodot}\ x{\isasyminv}{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isachardot}left{\isacharunderscore}unit{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{also}\ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isasymone}{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isachardot}left{\isacharunderscore}inverse{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{finally}\ \isamarkupfalse% +\isacommand{show}\ {\isacharquery}thesis\ \isamarkupfalse% +\isacommand{{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% +% +\begin{isamarkuptext}% +\noindent With \isa{group{\isacharunderscore}right{\isacharunderscore}inverse} already available, \isa{group{\isacharunderscore}right{\isacharunderscore}unit}\label{thm:group-right-unit} is now established + much easier.% +\end{isamarkuptext}% +\isamarkupfalse% +\isacommand{theorem}\ group{\isacharunderscore}right{\isacharunderscore}unit{\isacharcolon}\ {\isachardoublequote}x\ {\isasymodot}\ {\isasymone}\ {\isacharequal}\ {\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isasymColon}group{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkupfalse% +\isacommand{proof}\ {\isacharminus}\isanewline +\ \ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}x\ {\isasymodot}\ {\isasymone}\ {\isacharequal}\ x\ {\isasymodot}\ {\isacharparenleft}x{\isasyminv}\ {\isasymodot}\ x{\isacharparenright}{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isachardot}left{\isacharunderscore}inverse{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{also}\ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ x\ {\isasymodot}\ x{\isasyminv}\ {\isasymodot}\ x{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ semigroup{\isachardot}assoc{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{also}\ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isasymone}\ {\isasymodot}\ x{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isacharunderscore}right{\isacharunderscore}inverse{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{also}\ \isamarkupfalse% +\isacommand{have}\ {\isachardoublequote}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ x{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isachardot}left{\isacharunderscore}unit{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{finally}\ \isamarkupfalse% +\isacommand{show}\ {\isacharquery}thesis\ \isamarkupfalse% +\isacommand{{\isachardot}}\isanewline +\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% +% +\begin{isamarkuptext}% +\medskip Abstract theorems may be instantiated to only those types + \isa{{\isasymtau}} where the appropriate class membership \isa{{\isasymtau}\ {\isasymColon}\ c} is + known at Isabelle's type signature level. Since we have \isa{agroup\ {\isasymsubseteq}\ group\ {\isasymsubseteq}\ semigroup} by definition, all theorems of \isa{semigroup} and \isa{group} are automatically inherited by \isa{group} and \isa{agroup}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Abstract instantiation% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +From the definition, the \isa{monoid} and \isa{group} classes + have been independent. Note that for monoids, \isa{right{\isacharunderscore}unit} + had to be included as an axiom, but for groups both \isa{right{\isacharunderscore}unit} and \isa{right{\isacharunderscore}inverse} are derivable from the other + axioms. With \isa{group{\isacharunderscore}right{\isacharunderscore}unit} derived as a theorem of group + theory (see page~\pageref{thm:group-right-unit}), we may now + instantiate \isa{monoid\ {\isasymsubseteq}\ semigroup} and \isa{group\ {\isasymsubseteq}\ monoid} properly as follows (cf.\ \figref{fig:monoid-group}). + + \begin{figure}[htbp] + \begin{center} + \small + \unitlength 0.6mm + \begin{picture}(65,90)(0,-10) + \put(15,10){\line(0,1){10}} \put(15,30){\line(0,1){10}} + \put(15,50){\line(1,1){10}} \put(35,60){\line(1,-1){10}} + \put(15,5){\makebox(0,0){\isa{agroup}}} + \put(15,25){\makebox(0,0){\isa{group}}} + \put(15,45){\makebox(0,0){\isa{semigroup}}} + \put(30,65){\makebox(0,0){\isa{type}}} \put(50,45){\makebox(0,0){\isa{monoid}}} + \end{picture} + \hspace{4em} + \begin{picture}(30,90)(0,0) + \put(15,10){\line(0,1){10}} \put(15,30){\line(0,1){10}} + \put(15,50){\line(0,1){10}} \put(15,70){\line(0,1){10}} + \put(15,5){\makebox(0,0){\isa{agroup}}} + \put(15,25){\makebox(0,0){\isa{group}}} + \put(15,45){\makebox(0,0){\isa{monoid}}} + \put(15,65){\makebox(0,0){\isa{semigroup}}} + \put(15,85){\makebox(0,0){\isa{type}}} + \end{picture} + \caption{Monoids and groups: according to definition, and by proof} + \label{fig:monoid-group} + \end{center} + \end{figure}% +\end{isamarkuptext}% +\isamarkupfalse% +\isacommand{instance}\ monoid\ {\isasymsubseteq}\ semigroup\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{fix}\ x\ y\ z\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a{\isasymColon}monoid{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}x\ {\isasymodot}\ y\ {\isasymodot}\ z\ {\isacharequal}\ x\ {\isasymodot}\ {\isacharparenleft}y\ {\isasymodot}\ z{\isacharparenright}{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}rule\ monoid{\isachardot}assoc{\isacharparenright}\isanewline +\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\isamarkupfalse% +\isacommand{instance}\ group\ {\isasymsubseteq}\ monoid\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkupfalse% +\isacommand{proof}\isanewline +\ \ \isamarkupfalse% +\isacommand{fix}\ x\ y\ z\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a{\isasymColon}group{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}x\ {\isasymodot}\ y\ {\isasymodot}\ z\ {\isacharequal}\ x\ {\isasymodot}\ {\isacharparenleft}y\ {\isasymodot}\ z{\isacharparenright}{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}rule\ semigroup{\isachardot}assoc{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}{\isasymone}\ {\isasymodot}\ x\ {\isacharequal}\ x{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}rule\ group{\isachardot}left{\isacharunderscore}unit{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}x\ {\isasymodot}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}rule\ group{\isacharunderscore}right{\isacharunderscore}unit{\isacharparenright}\isanewline +\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% +% +\begin{isamarkuptext}% +\medskip The $\INSTANCE$ command sets up an appropriate goal that + represents the class inclusion (or type arity, see + \secref{sec:inst-arity}) to be proven (see also + \cite{isabelle-isar-ref}). The initial proof step causes + back-chaining of class membership statements wrt.\ the hierarchy of + any classes defined in the current theory; the effect is to reduce + to the initial statement to a number of goals that directly + correspond to any class axioms encountered on the path upwards + through the class hierarchy.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Concrete instantiation \label{sec:inst-arity}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +So far we have covered the case of the form $\INSTANCE$~\isa{c\isactrlsub {\isadigit{1}}\ {\isasymsubseteq}\ c\isactrlsub {\isadigit{2}}}, namely \emph{abstract instantiation} --- + $c@1$ is more special than \isa{c\isactrlsub {\isadigit{1}}} and thus an instance + of \isa{c\isactrlsub {\isadigit{2}}}. Even more interesting for practical + applications are \emph{concrete instantiations} of axiomatic type + classes. That is, certain simple schemes \isa{{\isacharparenleft}{\isasymalpha}\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymalpha}\isactrlsub n{\isacharparenright}\ t\ {\isasymColon}\ c} of class membership may be established at the + logical level and then transferred to Isabelle's type signature + level. + + \medskip As a typical example, we show that type \isa{bool} with + exclusive-or as \isa{{\isasymodot}} operation, identity as \isa{{\isasyminv}}, and + \isa{False} as \isa{{\isasymone}} forms an Abelian group.% +\end{isamarkuptext}% +\isamarkupfalse% +\isacommand{defs}\ {\isacharparenleft}\isakeyword{overloaded}{\isacharparenright}\isanewline +\ \ times{\isacharunderscore}bool{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequote}x\ {\isasymodot}\ y\ {\isasymequiv}\ x\ {\isasymnoteq}\ {\isacharparenleft}y{\isasymColon}bool{\isacharparenright}{\isachardoublequote}\isanewline +\ \ inverse{\isacharunderscore}bool{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequote}x{\isasyminv}\ {\isasymequiv}\ x{\isasymColon}bool{\isachardoublequote}\isanewline +\ \ unit{\isacharunderscore}bool{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequote}{\isasymone}\ {\isasymequiv}\ False{\isachardoublequote}\isamarkuptrue% +% +\begin{isamarkuptext}% +\medskip It is important to note that above $\DEFS$ are just + overloaded meta-level constant definitions, where type classes are + not yet involved at all. This form of constant definition with + overloading (and optional recursion over the syntactic structure of + simple types) are admissible as definitional extensions of plain HOL + \cite{Wenzel:1997:TPHOL}. The Haskell-style type system is not + required for overloading. Nevertheless, overloaded definitions are + best applied in the context of type classes. + + \medskip Since we have chosen above $\DEFS$ of the generic group + operations on type \isa{bool} appropriately, the class membership + \isa{bool\ {\isasymColon}\ agroup} may be now derived as follows.% +\end{isamarkuptext}% +\isamarkupfalse% +\isacommand{instance}\ bool\ {\isacharcolon}{\isacharcolon}\ agroup\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkupfalse% +\isacommand{proof}\ {\isacharparenleft}intro{\isacharunderscore}classes{\isacharcomma}\isanewline +\ \ \ \ unfold\ times{\isacharunderscore}bool{\isacharunderscore}def\ inverse{\isacharunderscore}bool{\isacharunderscore}def\ unit{\isacharunderscore}bool{\isacharunderscore}def{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{fix}\ x\ y\ z\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}{\isacharparenleft}{\isacharparenleft}x\ {\isasymnoteq}\ y{\isacharparenright}\ {\isasymnoteq}\ z{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isasymnoteq}\ {\isacharparenleft}y\ {\isasymnoteq}\ z{\isacharparenright}{\isacharparenright}{\isachardoublequote}\ \isamarkupfalse% +\isacommand{by}\ blast\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}{\isacharparenleft}False\ {\isasymnoteq}\ x{\isacharparenright}\ {\isacharequal}\ x{\isachardoublequote}\ \isamarkupfalse% +\isacommand{by}\ blast\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}{\isacharparenleft}x\ {\isasymnoteq}\ x{\isacharparenright}\ {\isacharequal}\ False{\isachardoublequote}\ \isamarkupfalse% +\isacommand{by}\ blast\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\ {\isachardoublequote}{\isacharparenleft}x\ {\isasymnoteq}\ y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}y\ {\isasymnoteq}\ x{\isacharparenright}{\isachardoublequote}\ \isamarkupfalse% +\isacommand{by}\ blast\isanewline +\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% +% +\begin{isamarkuptext}% +The result of an $\INSTANCE$ statement is both expressed as a + theorem of Isabelle's meta-logic, and as a type arity of the type + signature. The latter enables type-inference system to take care of + this new instance automatically. + + \medskip We could now also instantiate our group theory classes to + many other concrete types. For example, \isa{int\ {\isasymColon}\ agroup} + (e.g.\ by defining \isa{{\isasymodot}} as addition, \isa{{\isasyminv}} as negation + and \isa{{\isasymone}} as zero) or \isa{list\ {\isasymColon}\ {\isacharparenleft}type{\isacharparenright}\ semigroup} + (e.g.\ if \isa{{\isasymodot}} is defined as list append). Thus, the + characteristic constants \isa{{\isasymodot}}, \isa{{\isasyminv}}, \isa{{\isasymone}} + really become overloaded, i.e.\ have different meanings on different + types.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Lifting and Functors% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +As already mentioned above, overloading in the simply-typed HOL + systems may include recursion over the syntactic structure of types. + That is, definitional equations \isa{c\isactrlsup {\isasymtau}\ {\isasymequiv}\ t} may also + contain constants of name \isa{c} on the right-hand side --- if + these have types that are structurally simpler than \isa{{\isasymtau}}. + + This feature enables us to \emph{lift operations}, say to Cartesian + products, direct sums or function spaces. Subsequently we lift + \isa{{\isasymodot}} component-wise to binary products \isa{{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b}.% +\end{isamarkuptext}% +\isamarkupfalse% +\isacommand{defs}\ {\isacharparenleft}\isakeyword{overloaded}{\isacharparenright}\isanewline +\ \ times{\isacharunderscore}prod{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequote}p\ {\isasymodot}\ q\ {\isasymequiv}\ {\isacharparenleft}fst\ p\ {\isasymodot}\ fst\ q{\isacharcomma}\ snd\ p\ {\isasymodot}\ snd\ q{\isacharparenright}{\isachardoublequote}\isamarkuptrue% +% +\begin{isamarkuptext}% +It is very easy to see that associativity of \isa{{\isasymodot}} on \isa{{\isacharprime}a} + and \isa{{\isasymodot}} on \isa{{\isacharprime}b} transfers to \isa{{\isasymodot}} on \isa{{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b}. Hence the binary type constructor \isa{{\isasymodot}} maps semigroups + to semigroups. This may be established formally as follows.% +\end{isamarkuptext}% +\isamarkupfalse% +\isacommand{instance}\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}semigroup{\isacharcomma}\ semigroup{\isacharparenright}\ semigroup\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkupfalse% +\isacommand{proof}\ {\isacharparenleft}intro{\isacharunderscore}classes{\isacharcomma}\ unfold\ times{\isacharunderscore}prod{\isacharunderscore}def{\isacharparenright}\isanewline +\ \ \isamarkupfalse% +\isacommand{fix}\ p\ q\ r\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a{\isasymColon}semigroup\ {\isasymtimes}\ {\isacharprime}b{\isasymColon}semigroup{\isachardoublequote}\isanewline +\ \ \isamarkupfalse% +\isacommand{show}\isanewline +\ \ \ \ {\isachardoublequote}{\isacharparenleft}fst\ {\isacharparenleft}fst\ p\ {\isasymodot}\ fst\ q{\isacharcomma}\ snd\ p\ {\isasymodot}\ snd\ q{\isacharparenright}\ {\isasymodot}\ fst\ r{\isacharcomma}\isanewline +\ \ \ \ \ \ snd\ {\isacharparenleft}fst\ p\ {\isasymodot}\ fst\ q{\isacharcomma}\ snd\ p\ {\isasymodot}\ snd\ q{\isacharparenright}\ {\isasymodot}\ snd\ r{\isacharparenright}\ {\isacharequal}\isanewline +\ \ \ \ \ \ \ {\isacharparenleft}fst\ p\ {\isasymodot}\ fst\ {\isacharparenleft}fst\ q\ {\isasymodot}\ fst\ r{\isacharcomma}\ snd\ q\ {\isasymodot}\ snd\ r{\isacharparenright}{\isacharcomma}\isanewline +\ \ \ \ \ \ \ \ snd\ p\ {\isasymodot}\ snd\ {\isacharparenleft}fst\ q\ {\isasymodot}\ fst\ r{\isacharcomma}\ snd\ q\ {\isasymodot}\ snd\ r{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isanewline +\ \ \ \ \isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}simp\ add{\isacharcolon}\ semigroup{\isachardot}assoc{\isacharparenright}\isanewline +\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% +% +\begin{isamarkuptext}% +Thus, if we view class instances as ``structures'', then overloaded + constant definitions with recursion over types indirectly provide + some kind of ``functors'' --- i.e.\ mappings between abstract + theories.% +\end{isamarkuptext}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isamarkupfalse% +\isacommand{end}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 153fe83804c9 -r 096792bdc58e doc-src/AxClass/Group/document/Product.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/AxClass/Group/document/Product.tex Fri Aug 19 22:44:01 2005 +0200 @@ -0,0 +1,132 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Product}% +\isamarkuptrue% +% +\isamarkupheader{Syntactic classes% +} +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isamarkupfalse% +\isacommand{theory}\ Product\ \isakeyword{imports}\ Main\ \isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% +% +\begin{isamarkuptext}% +\medskip\noindent There is still a feature of Isabelle's type system + left that we have not yet discussed. When declaring polymorphic + constants \isa{c\ {\isasymColon}\ {\isasymsigma}}, the type variables occurring in \isa{{\isasymsigma}} + may be constrained by type classes (or even general sorts) in an + arbitrary way. Note that by default, in Isabelle/HOL the + declaration \isa{{\isasymodot}\ {\isasymColon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a} is actually an abbreviation + for \isa{{\isasymodot}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}type\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a} Since class \isa{type} is the + universal class of HOL, this is not really a constraint at all. + + The \isa{product} class below provides a less degenerate example of + syntactic type classes.% +\end{isamarkuptext}% +\isamarkupfalse% +\isacommand{axclass}\isanewline +\ \ product\ {\isasymsubseteq}\ type\isanewline +\isamarkupfalse% +\isacommand{consts}\isanewline +\ \ product\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a{\isasymColon}product\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isasymodot}{\isachardoublequote}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isamarkuptrue% +% +\begin{isamarkuptext}% +Here class \isa{product} is defined as subclass of \isa{type} + without any additional axioms. This effects in logical equivalence + of \isa{product} and \isa{type}, as is reflected by the trivial + introduction rule generated for this definition. + + \medskip So what is the difference of declaring \isa{{\isasymodot}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}product\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a} vs.\ declaring \isa{{\isasymodot}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}type\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a} anyway? In this particular case where \isa{product\ {\isasymequiv}\ type}, it should be obvious that both declarations are the same + from the logic's point of view. It even makes the most sense to + remove sort constraints from constant declarations, as far as the + purely logical meaning is concerned \cite{Wenzel:1997:TPHOL}. + + On the other hand there are syntactic differences, of course. + Constants \isa{{\isasymodot}} on some type \isa{{\isasymtau}} are rejected by the + type-checker, unless the arity \isa{{\isasymtau}\ {\isasymColon}\ product} is part of the + type signature. In our example, this arity may be always added when + required by means of an $\INSTANCE$ with the default proof $\DDOT$. + + \medskip Thus, we may observe the following discipline of using + syntactic classes. Overloaded polymorphic constants have their type + arguments restricted to an associated (logically trivial) class + \isa{c}. Only immediately before \emph{specifying} these + constants on a certain type \isa{{\isasymtau}} do we instantiate \isa{{\isasymtau}\ {\isasymColon}\ c}. + + This is done for class \isa{product} and type \isa{bool} as + follows.% +\end{isamarkuptext}% +\isamarkupfalse% +\isacommand{instance}\ bool\ {\isacharcolon}{\isacharcolon}\ product% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isamarkupfalse% +\isacommand{{\isachardot}{\isachardot}}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\isamarkupfalse% +\isacommand{defs}\ {\isacharparenleft}\isakeyword{overloaded}{\isacharparenright}\isanewline +\ \ product{\isacharunderscore}bool{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequote}x\ {\isasymodot}\ y\ {\isasymequiv}\ x\ {\isasymand}\ y{\isachardoublequote}\isamarkuptrue% +% +\begin{isamarkuptext}% +The definition \isa{prod{\isacharunderscore}bool{\isacharunderscore}def} becomes syntactically + well-formed only after the arity \isa{bool\ {\isasymColon}\ product} is made + known to the type checker. + + \medskip It is very important to see that above $\DEFS$ are not + directly connected with $\INSTANCE$ at all! We were just following + our convention to specify \isa{{\isasymodot}} on \isa{bool} after having + instantiated \isa{bool\ {\isasymColon}\ product}. Isabelle does not require + these definitions, which is in contrast to programming languages like + Haskell \cite{haskell-report}. + + \medskip While Isabelle type classes and those of Haskell are almost + the same as far as type-checking and type inference are concerned, + there are important semantic differences. Haskell classes require + their instances to \emph{provide operations} of certain \emph{names}. + Therefore, its \texttt{instance} has a \texttt{where} part that tells + the system what these ``member functions'' should be. + + This style of \texttt{instance} would not make much sense in + Isabelle's meta-logic, because there is no internal notion of + ``providing operations'' or even ``names of functions''.% +\end{isamarkuptext}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isamarkupfalse% +\isacommand{end}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 153fe83804c9 -r 096792bdc58e doc-src/AxClass/Group/document/Semigroups.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/AxClass/Group/document/Semigroups.tex Fri Aug 19 22:44:01 2005 +0200 @@ -0,0 +1,88 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Semigroups}% +\isamarkuptrue% +% +\isamarkupheader{Semigroups% +} +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isamarkupfalse% +\isacommand{theory}\ Semigroups\ \isakeyword{imports}\ Main\ \isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% +% +\begin{isamarkuptext}% +\medskip\noindent An axiomatic type class is simply a class of types + that all meet certain properties, which are also called \emph{class + axioms}. Thus, type classes may be also understood as type + predicates --- i.e.\ abstractions over a single type argument \isa{{\isacharprime}a}. Class axioms typically contain polymorphic constants that + depend on this type \isa{{\isacharprime}a}. These \emph{characteristic + constants} behave like operations associated with the ``carrier'' + type \isa{{\isacharprime}a}. + + We illustrate these basic concepts by the following formulation of + semigroups.% +\end{isamarkuptext}% +\isamarkupfalse% +\isacommand{consts}\isanewline +\ \ times\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isasymodot}{\isachardoublequote}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline +\isamarkupfalse% +\isacommand{axclass}\ semigroup\ {\isasymsubseteq}\ type\isanewline +\ \ assoc{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x\ {\isasymodot}\ y{\isacharparenright}\ {\isasymodot}\ z\ {\isacharequal}\ x\ {\isasymodot}\ {\isacharparenleft}y\ {\isasymodot}\ z{\isacharparenright}{\isachardoublequote}\isamarkuptrue% +% +\begin{isamarkuptext}% +\noindent Above we have first declared a polymorphic constant \isa{{\isasymodot}\ {\isasymColon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a} and then defined the class \isa{semigroup} of + all types \isa{{\isasymtau}} such that \isa{{\isasymodot}\ {\isasymColon}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}} is indeed an + associative operator. The \isa{assoc} axiom contains exactly one + type variable, which is invisible in the above presentation, though. + Also note that free term variables (like \isa{x}, \isa{y}, + \isa{z}) are allowed for user convenience --- conceptually all of + these are bound by outermost universal quantifiers. + + \medskip In general, type classes may be used to describe + \emph{structures} with exactly one carrier \isa{{\isacharprime}a} and a fixed + \emph{signature}. Different signatures require different classes. + Below, class \isa{plus{\isacharunderscore}semigroup} represents semigroups \isa{{\isacharparenleft}{\isasymtau}{\isacharcomma}\ {\isasymoplus}\isactrlsup {\isasymtau}{\isacharparenright}}, while the original \isa{semigroup} would + correspond to semigroups of the form \isa{{\isacharparenleft}{\isasymtau}{\isacharcomma}\ {\isasymodot}\isactrlsup {\isasymtau}{\isacharparenright}}.% +\end{isamarkuptext}% +\isamarkupfalse% +\isacommand{consts}\isanewline +\ \ plus\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isasymoplus}{\isachardoublequote}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline +\isamarkupfalse% +\isacommand{axclass}\ plus{\isacharunderscore}semigroup\ {\isasymsubseteq}\ type\isanewline +\ \ assoc{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x\ {\isasymoplus}\ y{\isacharparenright}\ {\isasymoplus}\ z\ {\isacharequal}\ x\ {\isasymoplus}\ {\isacharparenleft}y\ {\isasymoplus}\ z{\isacharparenright}{\isachardoublequote}\isamarkuptrue% +% +\begin{isamarkuptext}% +\noindent Even if classes \isa{plus{\isacharunderscore}semigroup} and \isa{semigroup} both represent semigroups in a sense, they are certainly + not quite the same.% +\end{isamarkuptext}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isamarkupfalse% +\isacommand{end}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 153fe83804c9 -r 096792bdc58e doc-src/AxClass/Group/document/isabelle.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/AxClass/Group/document/isabelle.sty Fri Aug 19 22:44:01 2005 +0200 @@ -0,0 +1,191 @@ +%% +%% Author: Markus Wenzel, TU Muenchen +%% +%% macros for Isabelle generated LaTeX output +%% +%% $Id$ + +%%% Simple document preparation (based on theory token language and symbols) + +% isabelle environments + +\newcommand{\isabellecontext}{UNKNOWN} + +\newcommand{\isastyle}{\small\tt\slshape} +\newcommand{\isastyleminor}{\small\tt\slshape} +\newcommand{\isastylescript}{\footnotesize\tt\slshape} +\newcommand{\isastyletext}{\normalsize\rm} +\newcommand{\isastyletxt}{\rm} +\newcommand{\isastylecmt}{\rm} + +%symbol markup -- \emph achieves decent spacing via italic corrections +\newcommand{\isamath}[1]{\emph{$#1$}} +\newcommand{\isatext}[1]{\emph{#1}} +\newcommand{\isascriptstyle}{\def\isamath##1{##1}\def\isatext##1{\mbox{\isastylescript##1}}} +\newcommand{\isactrlsub}[1]{\emph{${}\sb{\mbox{\isastylescript{#1}}}$}} +\newcommand{\isactrlsup}[1]{\emph{${}\sp{\mbox{\isastylescript{#1}}}$}} +\newcommand{\isactrlisub}[1]{\emph{${}\sb{\mbox{\isastylescript{#1}}}$}} +\newcommand{\isactrlisup}[1]{\emph{${}\sp{\mbox{\isastylescript{#1}}}$}} +\newcommand{\isactrlbsub}{\emph\bgroup\begin{math}{}\sb\bgroup\mbox\bgroup\isastylescript} +\newcommand{\isactrlesub}{\egroup\egroup\end{math}\egroup} +\newcommand{\isactrlbsup}{\emph\bgroup\begin{math}{}\sp\bgroup\mbox\bgroup\isastylescript} +\newcommand{\isactrlesup}{\egroup\egroup\end{math}\egroup} +\newcommand{\isactrlbold}[1]{{\bfseries\upshape\boldmath#1}} + + +\newdimen\isa@parindent\newdimen\isa@parskip + +\newenvironment{isabellebody}{% +\isamarkuptrue\par% +\isa@parindent\parindent\parindent0pt% +\isa@parskip\parskip\parskip0pt% +\isastyle}{\par} + +\newenvironment{isabelle} +{\begin{trivlist}\begin{isabellebody}\item\relax} +{\end{isabellebody}\end{trivlist}} + +\newcommand{\isa}[1]{\emph{\isastyleminor #1}} + +\newcommand{\isaindent}[1]{\hphantom{#1}} +\newcommand{\isanewline}{\mbox{}\par\mbox{}} +\newcommand{\isasep}{} % override with e.g. \renewcommand{\isasep}{\vspace{1ex}} +\newcommand{\isadigit}[1]{#1} + +\chardef\isacharbang=`\! +\chardef\isachardoublequote=`\" +\chardef\isacharhash=`\# +\chardef\isachardollar=`\$ +\chardef\isacharpercent=`\% +\chardef\isacharampersand=`\& +\chardef\isacharprime=`\' +\chardef\isacharparenleft=`\( +\chardef\isacharparenright=`\) +\chardef\isacharasterisk=`\* +\chardef\isacharplus=`\+ +\chardef\isacharcomma=`\, +\chardef\isacharminus=`\- +\chardef\isachardot=`\. +\chardef\isacharslash=`\/ +\chardef\isacharcolon=`\: +\chardef\isacharsemicolon=`\; +\chardef\isacharless=`\< +\chardef\isacharequal=`\= +\chardef\isachargreater=`\> +\chardef\isacharquery=`\? +\chardef\isacharat=`\@ +\chardef\isacharbrackleft=`\[ +\chardef\isacharbackslash=`\\ +\chardef\isacharbrackright=`\] +\chardef\isacharcircum=`\^ +\chardef\isacharunderscore=`\_ +\chardef\isacharbackquote=`\` +\chardef\isacharbraceleft=`\{ +\chardef\isacharbar=`\| +\chardef\isacharbraceright=`\} +\chardef\isachartilde=`\~ + + +% keyword and section markup + +\newcommand{\isakeywordcharunderscore}{\_} +\newcommand{\isakeyword}[1] +{\emph{\bf\def\isachardot{.}\def\isacharunderscore{\isakeywordcharunderscore}% +\def\isacharbraceleft{\{}\def\isacharbraceright{\}}#1}} +\newcommand{\isacommand}[1]{\isakeyword{#1}} + +\newcommand{\isamarkupheader}[1]{\section{#1}} +\newcommand{\isamarkupchapter}[1]{\chapter{#1}} +\newcommand{\isamarkupsection}[1]{\section{#1}} +\newcommand{\isamarkupsubsection}[1]{\subsection{#1}} +\newcommand{\isamarkupsubsubsection}[1]{\subsubsection{#1}} +\newcommand{\isamarkupsect}[1]{\section{#1}} +\newcommand{\isamarkupsubsect}[1]{\subsection{#1}} +\newcommand{\isamarkupsubsubsect}[1]{\subsubsection{#1}} + +\newif\ifisamarkup +\newcommand{\isabeginpar}{\par\ifisamarkup\relax\else\medskip\fi} +\newcommand{\isaendpar}{\par\medskip} +\newenvironment{isapar}{\parindent\isa@parindent\parskip\isa@parskip\isabeginpar}{\isaendpar} +\newenvironment{isamarkuptext}{\isastyletext\begin{isapar}}{\end{isapar}} +\newenvironment{isamarkuptxt}{\isastyletxt\begin{isapar}}{\end{isapar}} +\newcommand{\isamarkupcmt}[1]{{\isastylecmt--- #1}} + + +% alternative styles + +\newcommand{\isabellestyle}{} +\def\isabellestyle#1{\csname isabellestyle#1\endcsname} + +\newcommand{\isabellestylett}{% +\renewcommand{\isastyle}{\small\tt}% +\renewcommand{\isastyleminor}{\small\tt}% +\renewcommand{\isastylescript}{\footnotesize\tt}% +} +\newcommand{\isabellestyleit}{% +\renewcommand{\isastyle}{\small\it}% +\renewcommand{\isastyleminor}{\it}% +\renewcommand{\isastylescript}{\footnotesize\it}% +\renewcommand{\isakeywordcharunderscore}{\mbox{-}}% +\renewcommand{\isacharbang}{\isamath{!}}% +\renewcommand{\isachardoublequote}{}% +\renewcommand{\isacharhash}{\isamath{\#}}% +\renewcommand{\isachardollar}{\isamath{\$}}% +\renewcommand{\isacharpercent}{\isamath{\%}}% +\renewcommand{\isacharampersand}{\isamath{\&}}% +\renewcommand{\isacharprime}{\isamath{\mskip2mu{'}\mskip-2mu}}% +\renewcommand{\isacharparenleft}{\isamath{(}}% +\renewcommand{\isacharparenright}{\isamath{)}}% +\renewcommand{\isacharasterisk}{\isamath{*}}% +\renewcommand{\isacharplus}{\isamath{+}}% +\renewcommand{\isacharcomma}{\isamath{\mathord,}}% +\renewcommand{\isacharminus}{\isamath{-}}% +\renewcommand{\isachardot}{\isamath{\mathord.}}% +\renewcommand{\isacharslash}{\isamath{/}}% +\renewcommand{\isacharcolon}{\isamath{\mathord:}}% +\renewcommand{\isacharsemicolon}{\isamath{\mathord;}}% +\renewcommand{\isacharless}{\isamath{<}}% +\renewcommand{\isacharequal}{\isamath{=}}% +\renewcommand{\isachargreater}{\isamath{>}}% +\renewcommand{\isacharat}{\isamath{@}}% +\renewcommand{\isacharbrackleft}{\isamath{[}}% +\renewcommand{\isacharbackslash}{\isamath{\backslash}}% +\renewcommand{\isacharbrackright}{\isamath{]}}% +\renewcommand{\isacharunderscore}{\mbox{-}}% +\renewcommand{\isacharbraceleft}{\isamath{\{}}% +\renewcommand{\isacharbar}{\isamath{\mid}}% +\renewcommand{\isacharbraceright}{\isamath{\}}}% +\renewcommand{\isachartilde}{\isamath{{}\sp{\sim}}}% +} + +\newcommand{\isabellestylesl}{% +\isabellestyleit% +\renewcommand{\isastyle}{\small\sl}% +\renewcommand{\isastyleminor}{\sl}% +\renewcommand{\isastylescript}{\footnotesize\sl}% +} + + +% tagged regions + +%plain TeX version of comment package -- much faster! +\let\isafmtname\fmtname\def\fmtname{plain} +\usepackage{comment} +\let\fmtname\isafmtname + +\newcommand{\isafold}[1]{\emph{$\langle\mathord{\mathit{#1}}\rangle$}} + +\newcommand{\isakeeptag}[1]% +{\includecomment{isadelim#1}\includecomment{isatag#1}\csarg\def{isafold#1}{}} +\newcommand{\isadroptag}[1]% +{\excludecomment{isadelim#1}\excludecomment{isatag#1}\csarg\def{isafold#1}{}} +\newcommand{\isafoldtag}[1]% +{\includecomment{isadelim#1}\excludecomment{isatag#1}\csarg\def{isafold#1}{\isafold{#1}}} + +\isakeeptag{theory} +\isakeeptag{proof} +\isakeeptag{ML} +\isakeeptag{visible} +\isadroptag{invisible} + +\IfFileExists{isabelletags.sty}{\usepackage{isabelletags}}{} diff -r 153fe83804c9 -r 096792bdc58e doc-src/AxClass/Group/document/isabellesym.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/AxClass/Group/document/isabellesym.sty Fri Aug 19 22:44:01 2005 +0200 @@ -0,0 +1,362 @@ +%% +%% Author: Markus Wenzel, TU Muenchen +%% +%% definitions of standard Isabelle symbols +%% +%% $Id$ + +% symbol definitions + +\newcommand{\isasymzero}{\isamath{\mathbf{0}}} %requires amssymb +\newcommand{\isasymone}{\isamath{\mathbf{1}}} %requires amssymb +\newcommand{\isasymtwo}{\isamath{\mathbf{2}}} %requires amssymb +\newcommand{\isasymthree}{\isamath{\mathbf{3}}} %requires amssymb +\newcommand{\isasymfour}{\isamath{\mathbf{4}}} %requires amssymb +\newcommand{\isasymfive}{\isamath{\mathbf{5}}} %requires amssymb +\newcommand{\isasymsix}{\isamath{\mathbf{6}}} %requires amssymb +\newcommand{\isasymseven}{\isamath{\mathbf{7}}} %requires amssymb +\newcommand{\isasymeight}{\isamath{\mathbf{8}}} %requires amssymb +\newcommand{\isasymnine}{\isamath{\mathbf{9}}} %requires amssymb +\newcommand{\isasymA}{\isamath{\mathcal{A}}} +\newcommand{\isasymB}{\isamath{\mathcal{B}}} +\newcommand{\isasymC}{\isamath{\mathcal{C}}} +\newcommand{\isasymD}{\isamath{\mathcal{D}}} +\newcommand{\isasymE}{\isamath{\mathcal{E}}} +\newcommand{\isasymF}{\isamath{\mathcal{F}}} +\newcommand{\isasymG}{\isamath{\mathcal{G}}} +\newcommand{\isasymH}{\isamath{\mathcal{H}}} +\newcommand{\isasymI}{\isamath{\mathcal{I}}} +\newcommand{\isasymJ}{\isamath{\mathcal{J}}} +\newcommand{\isasymK}{\isamath{\mathcal{K}}} +\newcommand{\isasymL}{\isamath{\mathcal{L}}} +\newcommand{\isasymM}{\isamath{\mathcal{M}}} +\newcommand{\isasymN}{\isamath{\mathcal{N}}} +\newcommand{\isasymO}{\isamath{\mathcal{O}}} +\newcommand{\isasymP}{\isamath{\mathcal{P}}} +\newcommand{\isasymQ}{\isamath{\mathcal{Q}}} +\newcommand{\isasymR}{\isamath{\mathcal{R}}} +\newcommand{\isasymS}{\isamath{\mathcal{S}}} +\newcommand{\isasymT}{\isamath{\mathcal{T}}} +\newcommand{\isasymU}{\isamath{\mathcal{U}}} +\newcommand{\isasymV}{\isamath{\mathcal{V}}} +\newcommand{\isasymW}{\isamath{\mathcal{W}}} +\newcommand{\isasymX}{\isamath{\mathcal{X}}} +\newcommand{\isasymY}{\isamath{\mathcal{Y}}} +\newcommand{\isasymZ}{\isamath{\mathcal{Z}}} +\newcommand{\isasyma}{\isamath{\mathrm{a}}} +\newcommand{\isasymb}{\isamath{\mathrm{b}}} +\newcommand{\isasymc}{\isamath{\mathrm{c}}} +\newcommand{\isasymd}{\isamath{\mathrm{d}}} +\newcommand{\isasyme}{\isamath{\mathrm{e}}} +\newcommand{\isasymf}{\isamath{\mathrm{f}}} +\newcommand{\isasymg}{\isamath{\mathrm{g}}} +\newcommand{\isasymh}{\isamath{\mathrm{h}}} +\newcommand{\isasymi}{\isamath{\mathrm{i}}} +\newcommand{\isasymj}{\isamath{\mathrm{j}}} +\newcommand{\isasymk}{\isamath{\mathrm{k}}} +\newcommand{\isasyml}{\isamath{\mathrm{l}}} +\newcommand{\isasymm}{\isamath{\mathrm{m}}} +\newcommand{\isasymn}{\isamath{\mathrm{n}}} +\newcommand{\isasymo}{\isamath{\mathrm{o}}} +\newcommand{\isasymp}{\isamath{\mathrm{p}}} +\newcommand{\isasymq}{\isamath{\mathrm{q}}} +\newcommand{\isasymr}{\isamath{\mathrm{r}}} +\newcommand{\isasyms}{\isamath{\mathrm{s}}} +\newcommand{\isasymt}{\isamath{\mathrm{t}}} +\newcommand{\isasymu}{\isamath{\mathrm{u}}} +\newcommand{\isasymv}{\isamath{\mathrm{v}}} +\newcommand{\isasymw}{\isamath{\mathrm{w}}} +\newcommand{\isasymx}{\isamath{\mathrm{x}}} +\newcommand{\isasymy}{\isamath{\mathrm{y}}} +\newcommand{\isasymz}{\isamath{\mathrm{z}}} +\newcommand{\isasymAA}{\isamath{\mathfrak{A}}} %requires eufrak +\newcommand{\isasymBB}{\isamath{\mathfrak{B}}} %requires eufrak +\newcommand{\isasymCC}{\isamath{\mathfrak{C}}} %requires eufrak +\newcommand{\isasymDD}{\isamath{\mathfrak{D}}} %requires eufrak +\newcommand{\isasymEE}{\isamath{\mathfrak{E}}} %requires eufrak +\newcommand{\isasymFF}{\isamath{\mathfrak{F}}} %requires eufrak +\newcommand{\isasymGG}{\isamath{\mathfrak{G}}} %requires eufrak +\newcommand{\isasymHH}{\isamath{\mathfrak{H}}} %requires eufrak +\newcommand{\isasymII}{\isamath{\mathfrak{I}}} %requires eufrak +\newcommand{\isasymJJ}{\isamath{\mathfrak{J}}} %requires eufrak +\newcommand{\isasymKK}{\isamath{\mathfrak{K}}} %requires eufrak +\newcommand{\isasymLL}{\isamath{\mathfrak{L}}} %requires eufrak +\newcommand{\isasymMM}{\isamath{\mathfrak{M}}} %requires eufrak +\newcommand{\isasymNN}{\isamath{\mathfrak{N}}} %requires eufrak +\newcommand{\isasymOO}{\isamath{\mathfrak{O}}} %requires eufrak +\newcommand{\isasymPP}{\isamath{\mathfrak{P}}} %requires eufrak +\newcommand{\isasymQQ}{\isamath{\mathfrak{Q}}} %requires eufrak +\newcommand{\isasymRR}{\isamath{\mathfrak{R}}} %requires eufrak +\newcommand{\isasymSS}{\isamath{\mathfrak{S}}} %requires eufrak +\newcommand{\isasymTT}{\isamath{\mathfrak{T}}} %requires eufrak +\newcommand{\isasymUU}{\isamath{\mathfrak{U}}} %requires eufrak +\newcommand{\isasymVV}{\isamath{\mathfrak{V}}} %requires eufrak +\newcommand{\isasymWW}{\isamath{\mathfrak{W}}} %requires eufrak +\newcommand{\isasymXX}{\isamath{\mathfrak{X}}} %requires eufrak +\newcommand{\isasymYY}{\isamath{\mathfrak{Y}}} %requires eufrak +\newcommand{\isasymZZ}{\isamath{\mathfrak{Z}}} %requires eufrak +\newcommand{\isasymaa}{\isamath{\mathfrak{a}}} %requires eufrak +\newcommand{\isasymbb}{\isamath{\mathfrak{b}}} %requires eufrak +\newcommand{\isasymcc}{\isamath{\mathfrak{c}}} %requires eufrak +\newcommand{\isasymdd}{\isamath{\mathfrak{d}}} %requires eufrak +\newcommand{\isasymee}{\isamath{\mathfrak{e}}} %requires eufrak +\newcommand{\isasymff}{\isamath{\mathfrak{f}}} %requires eufrak +\newcommand{\isasymgg}{\isamath{\mathfrak{g}}} %requires eufrak +\newcommand{\isasymhh}{\isamath{\mathfrak{h}}} %requires eufrak +\newcommand{\isasymii}{\isamath{\mathfrak{i}}} %requires eufrak +\newcommand{\isasymjj}{\isamath{\mathfrak{j}}} %requires eufrak +\newcommand{\isasymkk}{\isamath{\mathfrak{k}}} %requires eufrak +\newcommand{\isasymll}{\isamath{\mathfrak{l}}} %requires eufrak +\newcommand{\isasymmm}{\isamath{\mathfrak{m}}} %requires eufrak +\newcommand{\isasymnn}{\isamath{\mathfrak{n}}} %requires eufrak +\newcommand{\isasymoo}{\isamath{\mathfrak{o}}} %requires eufrak +\newcommand{\isasympp}{\isamath{\mathfrak{p}}} %requires eufrak +\newcommand{\isasymqq}{\isamath{\mathfrak{q}}} %requires eufrak +\newcommand{\isasymrr}{\isamath{\mathfrak{r}}} %requires eufrak +\newcommand{\isasymss}{\isamath{\mathfrak{s}}} %requires eufrak +\newcommand{\isasymtt}{\isamath{\mathfrak{t}}} %requires eufrak +\newcommand{\isasymuu}{\isamath{\mathfrak{u}}} %requires eufrak +\newcommand{\isasymvv}{\isamath{\mathfrak{v}}} %requires eufrak +\newcommand{\isasymww}{\isamath{\mathfrak{w}}} %requires eufrak +\newcommand{\isasymxx}{\isamath{\mathfrak{x}}} %requires eufrak +\newcommand{\isasymyy}{\isamath{\mathfrak{y}}} %requires eufrak +\newcommand{\isasymzz}{\isamath{\mathfrak{z}}} %requires eufrak +\newcommand{\isasymalpha}{\isamath{\alpha}} +\newcommand{\isasymbeta}{\isamath{\beta}} +\newcommand{\isasymgamma}{\isamath{\gamma}} +\newcommand{\isasymdelta}{\isamath{\delta}} +\newcommand{\isasymepsilon}{\isamath{\varepsilon}} +\newcommand{\isasymzeta}{\isamath{\zeta}} +\newcommand{\isasymeta}{\isamath{\eta}} +\newcommand{\isasymtheta}{\isamath{\vartheta}} +\newcommand{\isasymiota}{\isamath{\iota}} +\newcommand{\isasymkappa}{\isamath{\kappa}} +\newcommand{\isasymlambda}{\isamath{\lambda}} +\newcommand{\isasymmu}{\isamath{\mu}} +\newcommand{\isasymnu}{\isamath{\nu}} +\newcommand{\isasymxi}{\isamath{\xi}} +\newcommand{\isasympi}{\isamath{\pi}} +\newcommand{\isasymrho}{\isamath{\varrho}} +\newcommand{\isasymsigma}{\isamath{\sigma}} +\newcommand{\isasymtau}{\isamath{\tau}} +\newcommand{\isasymupsilon}{\isamath{\upsilon}} +\newcommand{\isasymphi}{\isamath{\varphi}} +\newcommand{\isasymchi}{\isamath{\chi}} +\newcommand{\isasympsi}{\isamath{\psi}} +\newcommand{\isasymomega}{\isamath{\omega}} +\newcommand{\isasymGamma}{\isamath{\Gamma}} +\newcommand{\isasymDelta}{\isamath{\Delta}} +\newcommand{\isasymTheta}{\isamath{\Theta}} +\newcommand{\isasymLambda}{\isamath{\Lambda}} +\newcommand{\isasymXi}{\isamath{\Xi}} +\newcommand{\isasymPi}{\isamath{\Pi}} +\newcommand{\isasymSigma}{\isamath{\Sigma}} +\newcommand{\isasymUpsilon}{\isamath{\Upsilon}} +\newcommand{\isasymPhi}{\isamath{\Phi}} +\newcommand{\isasymPsi}{\isamath{\Psi}} +\newcommand{\isasymOmega}{\isamath{\Omega}} +\newcommand{\isasymbool}{\isamath{\mathrm{I}\mkern-3.8mu\mathrm{B}}} +\newcommand{\isasymcomplex}{\isamath{\mathrm{C}\mkern-15mu{\phantom{\mathrm{t}}\vrule}\mkern9mu}} +\newcommand{\isasymnat}{\isamath{\mathrm{I}\mkern-3.8mu\mathrm{N}}} +\newcommand{\isasymrat}{\isamath{\mathrm{Q}\mkern-16mu{\phantom{\mathrm{t}}\vrule}\mkern10mu}} +\newcommand{\isasymreal}{\isamath{\mathrm{I}\mkern-3.8mu\mathrm{R}}} +\newcommand{\isasymint}{\isamath{\mathsf{Z}\mkern-7.5mu\mathsf{Z}}} +\newcommand{\isasymleftarrow}{\isamath{\leftarrow}} +\newcommand{\isasymlongleftarrow}{\isamath{\longleftarrow}} +\newcommand{\isasymrightarrow}{\isamath{\rightarrow}} +\newcommand{\isasymlongrightarrow}{\isamath{\longrightarrow}} +\newcommand{\isasymLeftarrow}{\isamath{\Leftarrow}} +\newcommand{\isasymLongleftarrow}{\isamath{\Longleftarrow}} +\newcommand{\isasymRightarrow}{\isamath{\Rightarrow}} +\newcommand{\isasymLongrightarrow}{\isamath{\Longrightarrow}} +\newcommand{\isasymleftrightarrow}{\isamath{\leftrightarrow}} +\newcommand{\isasymlongleftrightarrow}{\isamath{\longleftrightarrow}} +\newcommand{\isasymLeftrightarrow}{\isamath{\Leftrightarrow}} +\newcommand{\isasymLongleftrightarrow}{\isamath{\Longleftrightarrow}} +\newcommand{\isasymmapsto}{\isamath{\mapsto}} +\newcommand{\isasymlongmapsto}{\isamath{\longmapsto}} +\newcommand{\isasymmidarrow}{\isamath{\relbar}} +\newcommand{\isasymMidarrow}{\isamath{\Relbar}} +\newcommand{\isasymhookleftarrow}{\isamath{\hookleftarrow}} +\newcommand{\isasymhookrightarrow}{\isamath{\hookrightarrow}} +\newcommand{\isasymleftharpoondown}{\isamath{\leftharpoondown}} +\newcommand{\isasymrightharpoondown}{\isamath{\rightharpoondown}} +\newcommand{\isasymleftharpoonup}{\isamath{\leftharpoonup}} +\newcommand{\isasymrightharpoonup}{\isamath{\rightharpoonup}} +\newcommand{\isasymrightleftharpoons}{\isamath{\rightleftharpoons}} +\newcommand{\isasymdownharpoonleft}{\isamath{\downharpoonleft}} %requires amssymb +\newcommand{\isasymdownharpoonright}{\isamath{\downharpoonright}} %requires amssymb +\newcommand{\isasymupharpoonleft}{\isamath{\upharpoonleft}} %requires amssymb +\newcommand{\isasymupharpoonright}{\isamath{\upharpoonright}} %requires amssymb +\newcommand{\isasymrestriction}{\isamath{\restriction}} %requires amssymb +\newcommand{\isasymleadsto}{\isamath{\leadsto}} %requires amssymb +\newcommand{\isasymup}{\isamath{\uparrow}} +\newcommand{\isasymUp}{\isamath{\Uparrow}} +\newcommand{\isasymdown}{\isamath{\downarrow}} +\newcommand{\isasymDown}{\isamath{\Downarrow}} +\newcommand{\isasymupdown}{\isamath{\updownarrow}} +\newcommand{\isasymUpdown}{\isamath{\Updownarrow}} +\newcommand{\isasymlangle}{\isamath{\langle}} +\newcommand{\isasymrangle}{\isamath{\rangle}} +\newcommand{\isasymlceil}{\isamath{\lceil}} +\newcommand{\isasymrceil}{\isamath{\rceil}} +\newcommand{\isasymlfloor}{\isamath{\lfloor}} +\newcommand{\isasymrfloor}{\isamath{\rfloor}} +\newcommand{\isasymlparr}{\isamath{\mathopen{(\mkern-3mu\mid}}} +\newcommand{\isasymrparr}{\isamath{\mathclose{\mid\mkern-3mu)}}} +\newcommand{\isasymlbrakk}{\isamath{\mathopen{\lbrack\mkern-3mu\lbrack}}} +\newcommand{\isasymrbrakk}{\isamath{\mathclose{\rbrack\mkern-3mu\rbrack}}} +\newcommand{\isasymlbrace}{\isamath{\mathopen{\lbrace\mkern-4.5mu\mid}}} +\newcommand{\isasymrbrace}{\isamath{\mathclose{\mid\mkern-4.5mu\rbrace}}} +\newcommand{\isasymguillemotleft}{\isatext{\flqq}} %requires babel +\newcommand{\isasymguillemotright}{\isatext{\frqq}} %requires babel +\newcommand{\isasymColon}{\isamath{\mathrel{::}}} +\newcommand{\isasymnot}{\isamath{\neg}} +\newcommand{\isasymbottom}{\isamath{\bot}} +\newcommand{\isasymtop}{\isamath{\top}} +\newcommand{\isasymand}{\isamath{\wedge}} +\newcommand{\isasymAnd}{\isamath{\bigwedge}} +\newcommand{\isasymor}{\isamath{\vee}} +\newcommand{\isasymOr}{\isamath{\bigvee}} +\newcommand{\isasymforall}{\isamath{\forall\,}} +\newcommand{\isasymexists}{\isamath{\exists\,}} +\newcommand{\isasymnexists}{\isamath{\nexists\,}} %requires amssymb +\newcommand{\isasymbox}{\isamath{\Box}} %requires amssymb +\newcommand{\isasymdiamond}{\isamath{\Diamond}} %requires amssymb +\newcommand{\isasymturnstile}{\isamath{\vdash}} +\newcommand{\isasymTurnstile}{\isamath{\models}} +\newcommand{\isasymtturnstile}{\isamath{\vdash\!\!\!\vdash}} +\newcommand{\isasymTTurnstile}{\isamath{\mid\!\models}} +\newcommand{\isasymstileturn}{\isamath{\dashv}} +\newcommand{\isasymsurd}{\isamath{\surd}} +\newcommand{\isasymle}{\isamath{\le}} +\newcommand{\isasymge}{\isamath{\ge}} +\newcommand{\isasymlless}{\isamath{\ll}} +\newcommand{\isasymggreater}{\isamath{\gg}} +\newcommand{\isasymlesssim}{\isamath{\lesssim}} %requires amssymb +\newcommand{\isasymgreatersim}{\isamath{\gtrsim}} %requires amssymb +\newcommand{\isasymlessapprox}{\isamath{\lessapprox}} %requires amssymb +\newcommand{\isasymgreaterapprox}{\isamath{\gtrapprox}} %requires amssymb +\newcommand{\isasymin}{\isamath{\in}} +\newcommand{\isasymnotin}{\isamath{\notin}} +\newcommand{\isasymsubset}{\isamath{\subset}} +\newcommand{\isasymsupset}{\isamath{\supset}} +\newcommand{\isasymsubseteq}{\isamath{\subseteq}} +\newcommand{\isasymsupseteq}{\isamath{\supseteq}} +\newcommand{\isasymsqsubset}{\isamath{\sqsubset}} %requires amssymb +\newcommand{\isasymsqsupset}{\isamath{\sqsupset}} %requires amssymb +\newcommand{\isasymsqsubseteq}{\isamath{\sqsubseteq}} +\newcommand{\isasymsqsupseteq}{\isamath{\sqsupseteq}} +\newcommand{\isasyminter}{\isamath{\cap}} +\newcommand{\isasymInter}{\isamath{\bigcap\,}} +\newcommand{\isasymunion}{\isamath{\cup}} +\newcommand{\isasymUnion}{\isamath{\bigcup\,}} +\newcommand{\isasymsqunion}{\isamath{\sqcup}} +\newcommand{\isasymSqunion}{\isamath{\bigsqcup\,}} +\newcommand{\isasymsqinter}{\isamath{\sqcap}} +\newcommand{\isasymSqinter}{\isamath{\bigsqcap\,}} %requires amsmath +\newcommand{\isasymuplus}{\isamath{\uplus}} +\newcommand{\isasymUplus}{\isamath{\biguplus\,}} +\newcommand{\isasymnoteq}{\isamath{\not=}} +\newcommand{\isasymsim}{\isamath{\sim}} +\newcommand{\isasymdoteq}{\isamath{\doteq}} +\newcommand{\isasymsimeq}{\isamath{\simeq}} +\newcommand{\isasymapprox}{\isamath{\approx}} +\newcommand{\isasymasymp}{\isamath{\asymp}} +\newcommand{\isasymcong}{\isamath{\cong}} +\newcommand{\isasymsmile}{\isamath{\smile}} +\newcommand{\isasymequiv}{\isamath{\equiv}} +\newcommand{\isasymfrown}{\isamath{\frown}} +\newcommand{\isasympropto}{\isamath{\propto}} +\newcommand{\isasymsome}{\isamath{\epsilon\,}} +\newcommand{\isasymJoin}{\isamath{\Join}} %requires amssymb +\newcommand{\isasymbowtie}{\isamath{\bowtie}} +\newcommand{\isasymprec}{\isamath{\prec}} +\newcommand{\isasymsucc}{\isamath{\succ}} +\newcommand{\isasympreceq}{\isamath{\preceq}} +\newcommand{\isasymsucceq}{\isamath{\succeq}} +\newcommand{\isasymparallel}{\isamath{\parallel}} +\newcommand{\isasymbar}{\isamath{\mid}} +\newcommand{\isasymplusminus}{\isamath{\pm}} +\newcommand{\isasymminusplus}{\isamath{\mp}} +\newcommand{\isasymtimes}{\isamath{\times}} +\newcommand{\isasymdiv}{\isamath{\div}} +\newcommand{\isasymcdot}{\isamath{\cdot}} +\newcommand{\isasymstar}{\isamath{\star}} +\newcommand{\isasymbullet}{\boldmath\isamath{\mathchoice{\displaystyle{\cdot}}{\textstyle{\cdot}}{\scriptstyle{\bullet}}{\scriptscriptstyle{\bullet}}}} +\newcommand{\isasymcirc}{\isamath{\circ}} +\newcommand{\isasymdagger}{\isamath{\dagger}} +\newcommand{\isasymddagger}{\isamath{\ddagger}} +\newcommand{\isasymlhd}{\isamath{\lhd}} %requires amssymb +\newcommand{\isasymrhd}{\isamath{\rhd}} %requires amssymb +\newcommand{\isasymunlhd}{\isamath{\unlhd}} %requires amssymb +\newcommand{\isasymunrhd}{\isamath{\unrhd}} %requires amssymb +\newcommand{\isasymtriangleleft}{\isamath{\triangleleft}} +\newcommand{\isasymtriangleright}{\isamath{\triangleright}} +\newcommand{\isasymtriangle}{\isamath{\triangle}} +\newcommand{\isasymtriangleq}{\isamath{\triangleq}} %requires amssymb +\newcommand{\isasymoplus}{\isamath{\oplus}} +\newcommand{\isasymOplus}{\isamath{\bigoplus\,}} +\newcommand{\isasymotimes}{\isamath{\otimes}} +\newcommand{\isasymOtimes}{\isamath{\bigotimes\,}} +\newcommand{\isasymodot}{\isamath{\odot}} +\newcommand{\isasymOdot}{\isamath{\bigodot\,}} +\newcommand{\isasymominus}{\isamath{\ominus}} +\newcommand{\isasymoslash}{\isamath{\oslash}} +\newcommand{\isasymdots}{\isamath{\dots}} +\newcommand{\isasymcdots}{\isamath{\cdots}} +\newcommand{\isasymSum}{\isamath{\sum\,}} +\newcommand{\isasymProd}{\isamath{\prod\,}} +\newcommand{\isasymCoprod}{\isamath{\coprod\,}} +\newcommand{\isasyminfinity}{\isamath{\infty}} +\newcommand{\isasymintegral}{\isamath{\int\,}} +\newcommand{\isasymointegral}{\isamath{\oint\,}} +\newcommand{\isasymclubsuit}{\isamath{\clubsuit}} +\newcommand{\isasymdiamondsuit}{\isamath{\diamondsuit}} +\newcommand{\isasymheartsuit}{\isamath{\heartsuit}} +\newcommand{\isasymspadesuit}{\isamath{\spadesuit}} +\newcommand{\isasymaleph}{\isamath{\aleph}} +\newcommand{\isasymemptyset}{\isamath{\emptyset}} +\newcommand{\isasymnabla}{\isamath{\nabla}} +\newcommand{\isasympartial}{\isamath{\partial}} +\newcommand{\isasymRe}{\isamath{\Re}} +\newcommand{\isasymIm}{\isamath{\Im}} +\newcommand{\isasymflat}{\isamath{\flat}} +\newcommand{\isasymnatural}{\isamath{\natural}} +\newcommand{\isasymsharp}{\isamath{\sharp}} +\newcommand{\isasymangle}{\isamath{\angle}} +\newcommand{\isasymcopyright}{\isatext{\rm\copyright}} +\newcommand{\isasymregistered}{\isatext{\rm\textregistered}} +\newcommand{\isasymhyphen}{\isatext{\rm-}} +\newcommand{\isasyminverse}{\isamath{{}^{-1}}} +\newcommand{\isasymonesuperior}{\isamath{\mathonesuperior}} %requires latin1 +\newcommand{\isasymonequarter}{\isatext{\rm\textonequarter}} %requires latin1 +\newcommand{\isasymtwosuperior}{\isamath{\mathtwosuperior}} %requires latin1 +\newcommand{\isasymonehalf}{\isatext{\rm\textonehalf}} %requires latin1 +\newcommand{\isasymthreesuperior}{\isamath{\maththreesuperior}} %requires latin1 +\newcommand{\isasymthreequarters}{\isatext{\rm\textthreequarters}} %requires latin1 +\newcommand{\isasymordfeminine}{\isatext{\rm\textordfeminine}} +\newcommand{\isasymordmasculine}{\isatext{\rm\textordmasculine}} +\newcommand{\isasymsection}{\isatext{\rm\S}} +\newcommand{\isasymparagraph}{\isatext{\rm\P}} +\newcommand{\isasymexclamdown}{\isatext{\rm\textexclamdown}} +\newcommand{\isasymquestiondown}{\isatext{\rm\textquestiondown}} +\newcommand{\isasymeuro}{\isatext{\textgreek{\euro}}} %requires greek babel +\newcommand{\isasympounds}{\isamath{\pounds}} +\newcommand{\isasymyen}{\isatext{\yen}} %requires amssymb +\newcommand{\isasymcent}{\isatext{\textcent}} %requires textcomp +\newcommand{\isasymcurrency}{\isatext{\textcurrency}} %requires textcomp +\newcommand{\isasymdegree}{\isatext{\rm\textdegree}} %requires latin1 +\newcommand{\isasymamalg}{\isamath{\amalg}} +\newcommand{\isasymmho}{\isamath{\mho}} %requires amssymb +\newcommand{\isasymlozenge}{\isamath{\lozenge}} %requires amssymb +\newcommand{\isasymwp}{\isamath{\wp}} +\newcommand{\isasymwrong}{\isamath{\wr}} +\newcommand{\isasymstruct}{\isamath{\diamond}} +\newcommand{\isasymacute}{\isatext{\'\relax}} +\newcommand{\isasymindex}{\isatext{\i}} +\newcommand{\isasymdieresis}{\isatext{\"\relax}} +\newcommand{\isasymcedilla}{\isatext{\c\relax}} +\newcommand{\isasymhungarumlaut}{\isatext{\H\relax}} +\newcommand{\isasymspacespace}{\isamath{~~}} +\newcommand{\isasymmodule}{\isamath{\langle}\isakeyword{module}\isamath{\rangle}} diff -r 153fe83804c9 -r 096792bdc58e doc-src/AxClass/IsaMakefile --- a/doc-src/AxClass/IsaMakefile Fri Aug 19 22:28:23 2005 +0200 +++ b/doc-src/AxClass/IsaMakefile Fri Aug 19 22:44:01 2005 +0200 @@ -13,7 +13,7 @@ SRC = $(ISABELLE_HOME)/src OUT = $(ISABELLE_OUTPUT) LOG = $(OUT)/log -USEDIR = $(ISATOOL) usedir -d "" -D ../generated +USEDIR = $(ISATOOL) usedir -d false -D document ## Group @@ -26,7 +26,7 @@ $(LOG)/HOL-Group.gz: $(OUT)/HOL Group/ROOT.ML Group/Group.thy \ Group/Product.thy Group/Semigroups.thy @$(USEDIR) $(OUT)/HOL Group - @rm -f generated/pdfsetup.sty generated/session.tex + @rm -f Group/document/pdfsetup.sty Group/document/session.tex ## Nat @@ -38,7 +38,7 @@ $(LOG)/FOL-Nat.gz: $(OUT)/FOL Nat/ROOT.ML Nat/NatClass.ML Nat/NatClass.thy @$(USEDIR) $(OUT)/FOL Nat - @rm -f generated/pdfsetup.sty generated/session.tex + @rm -f Nat/document/*.sty Nat/document/session.tex ## clean diff -r 153fe83804c9 -r 096792bdc58e doc-src/AxClass/Nat/document/NatClass.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/AxClass/Nat/document/NatClass.tex Fri Aug 19 22:44:01 2005 +0200 @@ -0,0 +1,96 @@ +% +\begin{isabellebody}% +\def\isabellecontext{NatClass}% +\isamarkuptrue% +% +\isamarkupheader{Defining natural numbers in FOL \label{sec:ex-natclass}% +} +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isamarkupfalse% +\isacommand{theory}\ NatClass\ \isakeyword{imports}\ FOL\ \isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% +% +\begin{isamarkuptext}% +\medskip\noindent Axiomatic type classes abstract over exactly one + type argument. Thus, any \emph{axiomatic} theory extension where each + axiom refers to at most one type variable, may be trivially turned + into a \emph{definitional} one. + + We illustrate this with the natural numbers in + Isabelle/FOL.\footnote{See also + \url{http://isabelle.in.tum.de/library/FOL/ex/NatClass.html}}% +\end{isamarkuptext}% +\isamarkupfalse% +\isacommand{consts}\isanewline +\ \ zero\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ \ \ \ {\isacharparenleft}{\isachardoublequote}{\isasymzero}{\isachardoublequote}{\isacharparenright}\isanewline +\ \ Suc\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\isanewline +\ \ rec\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\isanewline +\isanewline +\isamarkupfalse% +\isacommand{axclass}\ nat\ {\isasymsubseteq}\ {\isachardoublequote}term{\isachardoublequote}\isanewline +\ \ induct{\isacharcolon}\ {\isachardoublequote}P{\isacharparenleft}{\isasymzero}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isasymAnd}x{\isachardot}\ P{\isacharparenleft}x{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharparenleft}Suc{\isacharparenleft}x{\isacharparenright}{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharparenleft}n{\isacharparenright}{\isachardoublequote}\isanewline +\ \ Suc{\isacharunderscore}inject{\isacharcolon}\ {\isachardoublequote}Suc{\isacharparenleft}m{\isacharparenright}\ {\isacharequal}\ Suc{\isacharparenleft}n{\isacharparenright}\ {\isasymLongrightarrow}\ m\ {\isacharequal}\ n{\isachardoublequote}\isanewline +\ \ Suc{\isacharunderscore}neq{\isacharunderscore}{\isadigit{0}}{\isacharcolon}\ {\isachardoublequote}Suc{\isacharparenleft}m{\isacharparenright}\ {\isacharequal}\ {\isasymzero}\ {\isasymLongrightarrow}\ R{\isachardoublequote}\isanewline +\ \ rec{\isacharunderscore}{\isadigit{0}}{\isacharcolon}\ {\isachardoublequote}rec{\isacharparenleft}{\isasymzero}{\isacharcomma}\ a{\isacharcomma}\ f{\isacharparenright}\ {\isacharequal}\ a{\isachardoublequote}\isanewline +\ \ rec{\isacharunderscore}Suc{\isacharcolon}\ {\isachardoublequote}rec{\isacharparenleft}Suc{\isacharparenleft}m{\isacharparenright}{\isacharcomma}\ a{\isacharcomma}\ f{\isacharparenright}\ {\isacharequal}\ f{\isacharparenleft}m{\isacharcomma}\ rec{\isacharparenleft}m{\isacharcomma}\ a{\isacharcomma}\ f{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isanewline +\isanewline +\isamarkupfalse% +\isacommand{constdefs}\isanewline +\ \ add\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a{\isacharcolon}{\isacharcolon}nat\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isacharplus}{\isachardoublequote}\ {\isadigit{6}}{\isadigit{0}}{\isacharparenright}\isanewline +\ \ {\isachardoublequote}m\ {\isacharplus}\ n\ {\isasymequiv}\ rec{\isacharparenleft}m{\isacharcomma}\ n{\isacharcomma}\ {\isasymlambda}x\ y{\isachardot}\ Suc{\isacharparenleft}y{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkuptrue% +% +\begin{isamarkuptext}% +This is an abstract version of the plain \isa{Nat} theory in + FOL.\footnote{See + \url{http://isabelle.in.tum.de/library/FOL/ex/Nat.html}} Basically, + we have just replaced all occurrences of type \isa{nat} by \isa{{\isacharprime}a} and used the natural number axioms to define class \isa{nat}. + There is only a minor snag, that the original recursion operator + \isa{rec} had to be made monomorphic. + + Thus class \isa{nat} contains exactly those types \isa{{\isasymtau}} that + are isomorphic to ``the'' natural numbers (with signature \isa{{\isasymzero}}, \isa{Suc}, \isa{rec}). + + \medskip What we have done here can be also viewed as \emph{type + specification}. Of course, it still remains open if there is some + type at all that meets the class axioms. Now a very nice property of + axiomatic type classes is that abstract reasoning is always possible + --- independent of satisfiability. The meta-logic won't break, even + if some classes (or general sorts) turns out to be empty later --- + ``inconsistent'' class definitions may be useless, but do not cause + any harm. + + Theorems of the abstract natural numbers may be derived in the same + way as for the concrete version. The original proof scripts may be + re-used with some trivial changes only (mostly adding some type + constraints).% +\end{isamarkuptext}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isamarkupfalse% +\isacommand{end}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 153fe83804c9 -r 096792bdc58e doc-src/AxClass/axclass.tex --- a/doc-src/AxClass/axclass.tex Fri Aug 19 22:28:23 2005 +0200 +++ b/doc-src/AxClass/axclass.tex Fri Aug 19 22:44:01 2005 +0200 @@ -1,7 +1,7 @@ \documentclass[12pt,a4paper,fleqn]{report} \usepackage{graphicx,../iman,../extra,../isar} -\usepackage{generated/isabelle,generated/isabellesym} +\usepackage{Group/document/isabelle,Group/document/isabellesym} \usepackage{../pdfsetup} % last one! \isabellestyle{it} diff -r 153fe83804c9 -r 096792bdc58e doc-src/AxClass/body.tex --- a/doc-src/AxClass/body.tex Fri Aug 19 22:28:23 2005 +0200 +++ b/doc-src/AxClass/body.tex Fri Aug 19 22:44:01 2005 +0200 @@ -47,13 +47,13 @@ FOL. See also \url{http://isabelle.in.tum.de/library/HOL/AxClasses/} and \url{http://isabelle.in.tum.de/library/FOL/ex/NatClass.html}. -\input{generated/Semigroups} +\input{Group/document/Semigroups} -\input{generated/Group} +\input{Group/document/Group} -\input{generated/Product} +\input{Group/document/Product} -\input{generated/NatClass} +\input{Nat/document/NatClass} %% FIXME move some parts to ref or isar-ref manual (!?); diff -r 153fe83804c9 -r 096792bdc58e doc-src/AxClass/generated/Group.tex --- a/doc-src/AxClass/generated/Group.tex Fri Aug 19 22:28:23 2005 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,511 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Group}% -\isamarkuptrue% -% -\isamarkupheader{Basic group theory% -} -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isamarkupfalse% -\isacommand{theory}\ Group\ \isakeyword{imports}\ Main\ \isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isamarkuptrue% -% -\begin{isamarkuptext}% -\medskip\noindent The meta-level type system of Isabelle supports - \emph{intersections} and \emph{inclusions} of type classes. These - directly correspond to intersections and inclusions of type - predicates in a purely set theoretic sense. This is sufficient as a - means to describe simple hierarchies of structures. As an - illustration, we use the well-known example of semigroups, monoids, - general groups and Abelian groups.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Monoids and Groups% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -First we declare some polymorphic constants required later for the - signature parts of our structures.% -\end{isamarkuptext}% -\isamarkupfalse% -\isacommand{consts}\isanewline -\ \ times\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isasymodot}{\isachardoublequote}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline -\ \ invers\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\ \ \ \ {\isacharparenleft}{\isachardoublequote}{\isacharparenleft}{\isacharunderscore}{\isasyminv}{\isacharparenright}{\isachardoublequote}\ {\isacharbrackleft}{\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isadigit{0}}{\isacharbrackright}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharparenright}\isanewline -\ \ one\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ \ \ \ {\isacharparenleft}{\isachardoublequote}{\isasymone}{\isachardoublequote}{\isacharparenright}\isamarkuptrue% -% -\begin{isamarkuptext}% -\noindent Next we define class \isa{monoid} of monoids with - operations \isa{{\isasymodot}} and \isa{{\isasymone}}. Note that multiple class - axioms are allowed for user convenience --- they simply represent - the conjunction of their respective universal closures.% -\end{isamarkuptext}% -\isamarkupfalse% -\isacommand{axclass}\ monoid\ {\isasymsubseteq}\ type\isanewline -\ \ assoc{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x\ {\isasymodot}\ y{\isacharparenright}\ {\isasymodot}\ z\ {\isacharequal}\ x\ {\isasymodot}\ {\isacharparenleft}y\ {\isasymodot}\ z{\isacharparenright}{\isachardoublequote}\isanewline -\ \ left{\isacharunderscore}unit{\isacharcolon}\ {\isachardoublequote}{\isasymone}\ {\isasymodot}\ x\ {\isacharequal}\ x{\isachardoublequote}\isanewline -\ \ right{\isacharunderscore}unit{\isacharcolon}\ {\isachardoublequote}x\ {\isasymodot}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequote}\isamarkuptrue% -% -\begin{isamarkuptext}% -\noindent So class \isa{monoid} contains exactly those types - \isa{{\isasymtau}} where \isa{{\isasymodot}\ {\isasymColon}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}} and \isa{{\isasymone}\ {\isasymColon}\ {\isasymtau}} - are specified appropriately, such that \isa{{\isasymodot}} is associative and - \isa{{\isasymone}} is a left and right unit element for the \isa{{\isasymodot}} - operation.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -\medskip Independently of \isa{monoid}, we now define a linear - hierarchy of semigroups, general groups and Abelian groups. Note - that the names of class axioms are automatically qualified with each - class name, so we may re-use common names such as \isa{assoc}.% -\end{isamarkuptext}% -\isamarkupfalse% -\isacommand{axclass}\ semigroup\ {\isasymsubseteq}\ type\isanewline -\ \ assoc{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x\ {\isasymodot}\ y{\isacharparenright}\ {\isasymodot}\ z\ {\isacharequal}\ x\ {\isasymodot}\ {\isacharparenleft}y\ {\isasymodot}\ z{\isacharparenright}{\isachardoublequote}\isanewline -\isanewline -\isamarkupfalse% -\isacommand{axclass}\ group\ {\isasymsubseteq}\ semigroup\isanewline -\ \ left{\isacharunderscore}unit{\isacharcolon}\ {\isachardoublequote}{\isasymone}\ {\isasymodot}\ x\ {\isacharequal}\ x{\isachardoublequote}\isanewline -\ \ left{\isacharunderscore}inverse{\isacharcolon}\ {\isachardoublequote}x{\isasyminv}\ {\isasymodot}\ x\ {\isacharequal}\ {\isasymone}{\isachardoublequote}\isanewline -\isanewline -\isamarkupfalse% -\isacommand{axclass}\ agroup\ {\isasymsubseteq}\ group\isanewline -\ \ commute{\isacharcolon}\ {\isachardoublequote}x\ {\isasymodot}\ y\ {\isacharequal}\ y\ {\isasymodot}\ x{\isachardoublequote}\isamarkuptrue% -% -\begin{isamarkuptext}% -\noindent Class \isa{group} inherits associativity of \isa{{\isasymodot}} - from \isa{semigroup} and adds two further group axioms. Similarly, - \isa{agroup} is defined as the subset of \isa{group} such that - for all of its elements \isa{{\isasymtau}}, the operation \isa{{\isasymodot}\ {\isasymColon}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}} is even commutative.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Abstract reasoning% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In a sense, axiomatic type classes may be viewed as \emph{abstract - theories}. Above class definitions gives rise to abstract axioms - \isa{assoc}, \isa{left{\isacharunderscore}unit}, \isa{left{\isacharunderscore}inverse}, \isa{commute}, where any of these contain a type variable \isa{{\isacharprime}a\ {\isasymColon}\ c} that is restricted to types of the corresponding class \isa{c}. \emph{Sort constraints} like this express a logical - precondition for the whole formula. For example, \isa{assoc} - states that for all \isa{{\isasymtau}}, provided that \isa{{\isasymtau}\ {\isasymColon}\ semigroup}, the operation \isa{{\isasymodot}\ {\isasymColon}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}} is associative. - - \medskip From a technical point of view, abstract axioms are just - ordinary Isabelle theorems, which may be used in proofs without - special treatment. Such ``abstract proofs'' usually yield new - ``abstract theorems''. For example, we may now derive the following - well-known laws of general groups.% -\end{isamarkuptext}% -\isamarkupfalse% -\isacommand{theorem}\ group{\isacharunderscore}right{\isacharunderscore}inverse{\isacharcolon}\ {\isachardoublequote}x\ {\isasymodot}\ x{\isasyminv}\ {\isacharequal}\ {\isacharparenleft}{\isasymone}{\isasymColon}{\isacharprime}a{\isasymColon}group{\isacharparenright}{\isachardoublequote}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isamarkupfalse% -\isacommand{proof}\ {\isacharminus}\isanewline -\ \ \isamarkupfalse% -\isacommand{have}\ {\isachardoublequote}x\ {\isasymodot}\ x{\isasyminv}\ {\isacharequal}\ {\isasymone}\ {\isasymodot}\ {\isacharparenleft}x\ {\isasymodot}\ x{\isasyminv}{\isacharparenright}{\isachardoublequote}\isanewline -\ \ \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isachardot}left{\isacharunderscore}unit{\isacharparenright}\isanewline -\ \ \isamarkupfalse% -\isacommand{also}\ \isamarkupfalse% -\isacommand{have}\ {\isachardoublequote}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isasymone}\ {\isasymodot}\ x\ {\isasymodot}\ x{\isasyminv}{\isachardoublequote}\isanewline -\ \ \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ semigroup{\isachardot}assoc{\isacharparenright}\isanewline -\ \ \isamarkupfalse% -\isacommand{also}\ \isamarkupfalse% -\isacommand{have}\ {\isachardoublequote}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isacharparenleft}x{\isasyminv}{\isacharparenright}{\isasyminv}\ {\isasymodot}\ x{\isasyminv}\ {\isasymodot}\ x\ {\isasymodot}\ x{\isasyminv}{\isachardoublequote}\isanewline -\ \ \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isachardot}left{\isacharunderscore}inverse{\isacharparenright}\isanewline -\ \ \isamarkupfalse% -\isacommand{also}\ \isamarkupfalse% -\isacommand{have}\ {\isachardoublequote}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isacharparenleft}x{\isasyminv}{\isacharparenright}{\isasyminv}\ {\isasymodot}\ {\isacharparenleft}x{\isasyminv}\ {\isasymodot}\ x{\isacharparenright}\ {\isasymodot}\ x{\isasyminv}{\isachardoublequote}\isanewline -\ \ \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ semigroup{\isachardot}assoc{\isacharparenright}\isanewline -\ \ \isamarkupfalse% -\isacommand{also}\ \isamarkupfalse% -\isacommand{have}\ {\isachardoublequote}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isacharparenleft}x{\isasyminv}{\isacharparenright}{\isasyminv}\ {\isasymodot}\ {\isasymone}\ {\isasymodot}\ x{\isasyminv}{\isachardoublequote}\isanewline -\ \ \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isachardot}left{\isacharunderscore}inverse{\isacharparenright}\isanewline -\ \ \isamarkupfalse% -\isacommand{also}\ \isamarkupfalse% -\isacommand{have}\ {\isachardoublequote}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isacharparenleft}x{\isasyminv}{\isacharparenright}{\isasyminv}\ {\isasymodot}\ {\isacharparenleft}{\isasymone}\ {\isasymodot}\ x{\isasyminv}{\isacharparenright}{\isachardoublequote}\isanewline -\ \ \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ semigroup{\isachardot}assoc{\isacharparenright}\isanewline -\ \ \isamarkupfalse% -\isacommand{also}\ \isamarkupfalse% -\isacommand{have}\ {\isachardoublequote}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isacharparenleft}x{\isasyminv}{\isacharparenright}{\isasyminv}\ {\isasymodot}\ x{\isasyminv}{\isachardoublequote}\isanewline -\ \ \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isachardot}left{\isacharunderscore}unit{\isacharparenright}\isanewline -\ \ \isamarkupfalse% -\isacommand{also}\ \isamarkupfalse% -\isacommand{have}\ {\isachardoublequote}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isasymone}{\isachardoublequote}\isanewline -\ \ \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isachardot}left{\isacharunderscore}inverse{\isacharparenright}\isanewline -\ \ \isamarkupfalse% -\isacommand{finally}\ \isamarkupfalse% -\isacommand{show}\ {\isacharquery}thesis\ \isamarkupfalse% -\isacommand{{\isachardot}}\isanewline -\isamarkupfalse% -\isacommand{qed}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isamarkuptrue% -% -\begin{isamarkuptext}% -\noindent With \isa{group{\isacharunderscore}right{\isacharunderscore}inverse} already available, \isa{group{\isacharunderscore}right{\isacharunderscore}unit}\label{thm:group-right-unit} is now established - much easier.% -\end{isamarkuptext}% -\isamarkupfalse% -\isacommand{theorem}\ group{\isacharunderscore}right{\isacharunderscore}unit{\isacharcolon}\ {\isachardoublequote}x\ {\isasymodot}\ {\isasymone}\ {\isacharequal}\ {\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isasymColon}group{\isacharparenright}{\isachardoublequote}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isamarkupfalse% -\isacommand{proof}\ {\isacharminus}\isanewline -\ \ \isamarkupfalse% -\isacommand{have}\ {\isachardoublequote}x\ {\isasymodot}\ {\isasymone}\ {\isacharequal}\ x\ {\isasymodot}\ {\isacharparenleft}x{\isasyminv}\ {\isasymodot}\ x{\isacharparenright}{\isachardoublequote}\isanewline -\ \ \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isachardot}left{\isacharunderscore}inverse{\isacharparenright}\isanewline -\ \ \isamarkupfalse% -\isacommand{also}\ \isamarkupfalse% -\isacommand{have}\ {\isachardoublequote}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ x\ {\isasymodot}\ x{\isasyminv}\ {\isasymodot}\ x{\isachardoublequote}\isanewline -\ \ \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ semigroup{\isachardot}assoc{\isacharparenright}\isanewline -\ \ \isamarkupfalse% -\isacommand{also}\ \isamarkupfalse% -\isacommand{have}\ {\isachardoublequote}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ {\isasymone}\ {\isasymodot}\ x{\isachardoublequote}\isanewline -\ \ \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isacharunderscore}right{\isacharunderscore}inverse{\isacharparenright}\isanewline -\ \ \isamarkupfalse% -\isacommand{also}\ \isamarkupfalse% -\isacommand{have}\ {\isachardoublequote}{\isachardot}{\isachardot}{\isachardot}\ {\isacharequal}\ x{\isachardoublequote}\isanewline -\ \ \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ only{\isacharcolon}\ group{\isachardot}left{\isacharunderscore}unit{\isacharparenright}\isanewline -\ \ \isamarkupfalse% -\isacommand{finally}\ \isamarkupfalse% -\isacommand{show}\ {\isacharquery}thesis\ \isamarkupfalse% -\isacommand{{\isachardot}}\isanewline -\isamarkupfalse% -\isacommand{qed}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isamarkuptrue% -% -\begin{isamarkuptext}% -\medskip Abstract theorems may be instantiated to only those types - \isa{{\isasymtau}} where the appropriate class membership \isa{{\isasymtau}\ {\isasymColon}\ c} is - known at Isabelle's type signature level. Since we have \isa{agroup\ {\isasymsubseteq}\ group\ {\isasymsubseteq}\ semigroup} by definition, all theorems of \isa{semigroup} and \isa{group} are automatically inherited by \isa{group} and \isa{agroup}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Abstract instantiation% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -From the definition, the \isa{monoid} and \isa{group} classes - have been independent. Note that for monoids, \isa{right{\isacharunderscore}unit} - had to be included as an axiom, but for groups both \isa{right{\isacharunderscore}unit} and \isa{right{\isacharunderscore}inverse} are derivable from the other - axioms. With \isa{group{\isacharunderscore}right{\isacharunderscore}unit} derived as a theorem of group - theory (see page~\pageref{thm:group-right-unit}), we may now - instantiate \isa{monoid\ {\isasymsubseteq}\ semigroup} and \isa{group\ {\isasymsubseteq}\ monoid} properly as follows (cf.\ \figref{fig:monoid-group}). - - \begin{figure}[htbp] - \begin{center} - \small - \unitlength 0.6mm - \begin{picture}(65,90)(0,-10) - \put(15,10){\line(0,1){10}} \put(15,30){\line(0,1){10}} - \put(15,50){\line(1,1){10}} \put(35,60){\line(1,-1){10}} - \put(15,5){\makebox(0,0){\isa{agroup}}} - \put(15,25){\makebox(0,0){\isa{group}}} - \put(15,45){\makebox(0,0){\isa{semigroup}}} - \put(30,65){\makebox(0,0){\isa{type}}} \put(50,45){\makebox(0,0){\isa{monoid}}} - \end{picture} - \hspace{4em} - \begin{picture}(30,90)(0,0) - \put(15,10){\line(0,1){10}} \put(15,30){\line(0,1){10}} - \put(15,50){\line(0,1){10}} \put(15,70){\line(0,1){10}} - \put(15,5){\makebox(0,0){\isa{agroup}}} - \put(15,25){\makebox(0,0){\isa{group}}} - \put(15,45){\makebox(0,0){\isa{monoid}}} - \put(15,65){\makebox(0,0){\isa{semigroup}}} - \put(15,85){\makebox(0,0){\isa{type}}} - \end{picture} - \caption{Monoids and groups: according to definition, and by proof} - \label{fig:monoid-group} - \end{center} - \end{figure}% -\end{isamarkuptext}% -\isamarkupfalse% -\isacommand{instance}\ monoid\ {\isasymsubseteq}\ semigroup\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isamarkupfalse% -\isacommand{proof}\isanewline -\ \ \isamarkupfalse% -\isacommand{fix}\ x\ y\ z\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a{\isasymColon}monoid{\isachardoublequote}\isanewline -\ \ \isamarkupfalse% -\isacommand{show}\ {\isachardoublequote}x\ {\isasymodot}\ y\ {\isasymodot}\ z\ {\isacharequal}\ x\ {\isasymodot}\ {\isacharparenleft}y\ {\isasymodot}\ z{\isacharparenright}{\isachardoublequote}\isanewline -\ \ \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}rule\ monoid{\isachardot}assoc{\isacharparenright}\isanewline -\isamarkupfalse% -\isacommand{qed}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isamarkupfalse% -\isacommand{instance}\ group\ {\isasymsubseteq}\ monoid\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isamarkupfalse% -\isacommand{proof}\isanewline -\ \ \isamarkupfalse% -\isacommand{fix}\ x\ y\ z\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a{\isasymColon}group{\isachardoublequote}\isanewline -\ \ \isamarkupfalse% -\isacommand{show}\ {\isachardoublequote}x\ {\isasymodot}\ y\ {\isasymodot}\ z\ {\isacharequal}\ x\ {\isasymodot}\ {\isacharparenleft}y\ {\isasymodot}\ z{\isacharparenright}{\isachardoublequote}\isanewline -\ \ \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}rule\ semigroup{\isachardot}assoc{\isacharparenright}\isanewline -\ \ \isamarkupfalse% -\isacommand{show}\ {\isachardoublequote}{\isasymone}\ {\isasymodot}\ x\ {\isacharequal}\ x{\isachardoublequote}\isanewline -\ \ \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}rule\ group{\isachardot}left{\isacharunderscore}unit{\isacharparenright}\isanewline -\ \ \isamarkupfalse% -\isacommand{show}\ {\isachardoublequote}x\ {\isasymodot}\ {\isasymone}\ {\isacharequal}\ x{\isachardoublequote}\isanewline -\ \ \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}rule\ group{\isacharunderscore}right{\isacharunderscore}unit{\isacharparenright}\isanewline -\isamarkupfalse% -\isacommand{qed}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isamarkuptrue% -% -\begin{isamarkuptext}% -\medskip The $\INSTANCE$ command sets up an appropriate goal that - represents the class inclusion (or type arity, see - \secref{sec:inst-arity}) to be proven (see also - \cite{isabelle-isar-ref}). The initial proof step causes - back-chaining of class membership statements wrt.\ the hierarchy of - any classes defined in the current theory; the effect is to reduce - to the initial statement to a number of goals that directly - correspond to any class axioms encountered on the path upwards - through the class hierarchy.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Concrete instantiation \label{sec:inst-arity}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -So far we have covered the case of the form $\INSTANCE$~\isa{c\isactrlsub {\isadigit{1}}\ {\isasymsubseteq}\ c\isactrlsub {\isadigit{2}}}, namely \emph{abstract instantiation} --- - $c@1$ is more special than \isa{c\isactrlsub {\isadigit{1}}} and thus an instance - of \isa{c\isactrlsub {\isadigit{2}}}. Even more interesting for practical - applications are \emph{concrete instantiations} of axiomatic type - classes. That is, certain simple schemes \isa{{\isacharparenleft}{\isasymalpha}\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymalpha}\isactrlsub n{\isacharparenright}\ t\ {\isasymColon}\ c} of class membership may be established at the - logical level and then transferred to Isabelle's type signature - level. - - \medskip As a typical example, we show that type \isa{bool} with - exclusive-or as \isa{{\isasymodot}} operation, identity as \isa{{\isasyminv}}, and - \isa{False} as \isa{{\isasymone}} forms an Abelian group.% -\end{isamarkuptext}% -\isamarkupfalse% -\isacommand{defs}\ {\isacharparenleft}\isakeyword{overloaded}{\isacharparenright}\isanewline -\ \ times{\isacharunderscore}bool{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequote}x\ {\isasymodot}\ y\ {\isasymequiv}\ x\ {\isasymnoteq}\ {\isacharparenleft}y{\isasymColon}bool{\isacharparenright}{\isachardoublequote}\isanewline -\ \ inverse{\isacharunderscore}bool{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequote}x{\isasyminv}\ {\isasymequiv}\ x{\isasymColon}bool{\isachardoublequote}\isanewline -\ \ unit{\isacharunderscore}bool{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequote}{\isasymone}\ {\isasymequiv}\ False{\isachardoublequote}\isamarkuptrue% -% -\begin{isamarkuptext}% -\medskip It is important to note that above $\DEFS$ are just - overloaded meta-level constant definitions, where type classes are - not yet involved at all. This form of constant definition with - overloading (and optional recursion over the syntactic structure of - simple types) are admissible as definitional extensions of plain HOL - \cite{Wenzel:1997:TPHOL}. The Haskell-style type system is not - required for overloading. Nevertheless, overloaded definitions are - best applied in the context of type classes. - - \medskip Since we have chosen above $\DEFS$ of the generic group - operations on type \isa{bool} appropriately, the class membership - \isa{bool\ {\isasymColon}\ agroup} may be now derived as follows.% -\end{isamarkuptext}% -\isamarkupfalse% -\isacommand{instance}\ bool\ {\isacharcolon}{\isacharcolon}\ agroup\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isamarkupfalse% -\isacommand{proof}\ {\isacharparenleft}intro{\isacharunderscore}classes{\isacharcomma}\isanewline -\ \ \ \ unfold\ times{\isacharunderscore}bool{\isacharunderscore}def\ inverse{\isacharunderscore}bool{\isacharunderscore}def\ unit{\isacharunderscore}bool{\isacharunderscore}def{\isacharparenright}\isanewline -\ \ \isamarkupfalse% -\isacommand{fix}\ x\ y\ z\isanewline -\ \ \isamarkupfalse% -\isacommand{show}\ {\isachardoublequote}{\isacharparenleft}{\isacharparenleft}x\ {\isasymnoteq}\ y{\isacharparenright}\ {\isasymnoteq}\ z{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isasymnoteq}\ {\isacharparenleft}y\ {\isasymnoteq}\ z{\isacharparenright}{\isacharparenright}{\isachardoublequote}\ \isamarkupfalse% -\isacommand{by}\ blast\isanewline -\ \ \isamarkupfalse% -\isacommand{show}\ {\isachardoublequote}{\isacharparenleft}False\ {\isasymnoteq}\ x{\isacharparenright}\ {\isacharequal}\ x{\isachardoublequote}\ \isamarkupfalse% -\isacommand{by}\ blast\isanewline -\ \ \isamarkupfalse% -\isacommand{show}\ {\isachardoublequote}{\isacharparenleft}x\ {\isasymnoteq}\ x{\isacharparenright}\ {\isacharequal}\ False{\isachardoublequote}\ \isamarkupfalse% -\isacommand{by}\ blast\isanewline -\ \ \isamarkupfalse% -\isacommand{show}\ {\isachardoublequote}{\isacharparenleft}x\ {\isasymnoteq}\ y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}y\ {\isasymnoteq}\ x{\isacharparenright}{\isachardoublequote}\ \isamarkupfalse% -\isacommand{by}\ blast\isanewline -\isamarkupfalse% -\isacommand{qed}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isamarkuptrue% -% -\begin{isamarkuptext}% -The result of an $\INSTANCE$ statement is both expressed as a - theorem of Isabelle's meta-logic, and as a type arity of the type - signature. The latter enables type-inference system to take care of - this new instance automatically. - - \medskip We could now also instantiate our group theory classes to - many other concrete types. For example, \isa{int\ {\isasymColon}\ agroup} - (e.g.\ by defining \isa{{\isasymodot}} as addition, \isa{{\isasyminv}} as negation - and \isa{{\isasymone}} as zero) or \isa{list\ {\isasymColon}\ {\isacharparenleft}type{\isacharparenright}\ semigroup} - (e.g.\ if \isa{{\isasymodot}} is defined as list append). Thus, the - characteristic constants \isa{{\isasymodot}}, \isa{{\isasyminv}}, \isa{{\isasymone}} - really become overloaded, i.e.\ have different meanings on different - types.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Lifting and Functors% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -As already mentioned above, overloading in the simply-typed HOL - systems may include recursion over the syntactic structure of types. - That is, definitional equations \isa{c\isactrlsup {\isasymtau}\ {\isasymequiv}\ t} may also - contain constants of name \isa{c} on the right-hand side --- if - these have types that are structurally simpler than \isa{{\isasymtau}}. - - This feature enables us to \emph{lift operations}, say to Cartesian - products, direct sums or function spaces. Subsequently we lift - \isa{{\isasymodot}} component-wise to binary products \isa{{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b}.% -\end{isamarkuptext}% -\isamarkupfalse% -\isacommand{defs}\ {\isacharparenleft}\isakeyword{overloaded}{\isacharparenright}\isanewline -\ \ times{\isacharunderscore}prod{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequote}p\ {\isasymodot}\ q\ {\isasymequiv}\ {\isacharparenleft}fst\ p\ {\isasymodot}\ fst\ q{\isacharcomma}\ snd\ p\ {\isasymodot}\ snd\ q{\isacharparenright}{\isachardoublequote}\isamarkuptrue% -% -\begin{isamarkuptext}% -It is very easy to see that associativity of \isa{{\isasymodot}} on \isa{{\isacharprime}a} - and \isa{{\isasymodot}} on \isa{{\isacharprime}b} transfers to \isa{{\isasymodot}} on \isa{{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b}. Hence the binary type constructor \isa{{\isasymodot}} maps semigroups - to semigroups. This may be established formally as follows.% -\end{isamarkuptext}% -\isamarkupfalse% -\isacommand{instance}\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}semigroup{\isacharcomma}\ semigroup{\isacharparenright}\ semigroup\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isamarkupfalse% -\isacommand{proof}\ {\isacharparenleft}intro{\isacharunderscore}classes{\isacharcomma}\ unfold\ times{\isacharunderscore}prod{\isacharunderscore}def{\isacharparenright}\isanewline -\ \ \isamarkupfalse% -\isacommand{fix}\ p\ q\ r\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a{\isasymColon}semigroup\ {\isasymtimes}\ {\isacharprime}b{\isasymColon}semigroup{\isachardoublequote}\isanewline -\ \ \isamarkupfalse% -\isacommand{show}\isanewline -\ \ \ \ {\isachardoublequote}{\isacharparenleft}fst\ {\isacharparenleft}fst\ p\ {\isasymodot}\ fst\ q{\isacharcomma}\ snd\ p\ {\isasymodot}\ snd\ q{\isacharparenright}\ {\isasymodot}\ fst\ r{\isacharcomma}\isanewline -\ \ \ \ \ \ snd\ {\isacharparenleft}fst\ p\ {\isasymodot}\ fst\ q{\isacharcomma}\ snd\ p\ {\isasymodot}\ snd\ q{\isacharparenright}\ {\isasymodot}\ snd\ r{\isacharparenright}\ {\isacharequal}\isanewline -\ \ \ \ \ \ \ {\isacharparenleft}fst\ p\ {\isasymodot}\ fst\ {\isacharparenleft}fst\ q\ {\isasymodot}\ fst\ r{\isacharcomma}\ snd\ q\ {\isasymodot}\ snd\ r{\isacharparenright}{\isacharcomma}\isanewline -\ \ \ \ \ \ \ \ snd\ p\ {\isasymodot}\ snd\ {\isacharparenleft}fst\ q\ {\isasymodot}\ fst\ r{\isacharcomma}\ snd\ q\ {\isasymodot}\ snd\ r{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isanewline -\ \ \ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ add{\isacharcolon}\ semigroup{\isachardot}assoc{\isacharparenright}\isanewline -\isamarkupfalse% -\isacommand{qed}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isamarkuptrue% -% -\begin{isamarkuptext}% -Thus, if we view class instances as ``structures'', then overloaded - constant definitions with recursion over types indirectly provide - some kind of ``functors'' --- i.e.\ mappings between abstract - theories.% -\end{isamarkuptext}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isamarkupfalse% -\isacommand{end}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 153fe83804c9 -r 096792bdc58e doc-src/AxClass/generated/NatClass.tex --- a/doc-src/AxClass/generated/NatClass.tex Fri Aug 19 22:28:23 2005 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,96 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{NatClass}% -\isamarkuptrue% -% -\isamarkupheader{Defining natural numbers in FOL \label{sec:ex-natclass}% -} -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isamarkupfalse% -\isacommand{theory}\ NatClass\ \isakeyword{imports}\ FOL\ \isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isamarkuptrue% -% -\begin{isamarkuptext}% -\medskip\noindent Axiomatic type classes abstract over exactly one - type argument. Thus, any \emph{axiomatic} theory extension where each - axiom refers to at most one type variable, may be trivially turned - into a \emph{definitional} one. - - We illustrate this with the natural numbers in - Isabelle/FOL.\footnote{See also - \url{http://isabelle.in.tum.de/library/FOL/ex/NatClass.html}}% -\end{isamarkuptext}% -\isamarkupfalse% -\isacommand{consts}\isanewline -\ \ zero\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ \ \ \ {\isacharparenleft}{\isachardoublequote}{\isasymzero}{\isachardoublequote}{\isacharparenright}\isanewline -\ \ Suc\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\isanewline -\ \ rec\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\isanewline -\isanewline -\isamarkupfalse% -\isacommand{axclass}\ nat\ {\isasymsubseteq}\ {\isachardoublequote}term{\isachardoublequote}\isanewline -\ \ induct{\isacharcolon}\ {\isachardoublequote}P{\isacharparenleft}{\isasymzero}{\isacharparenright}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isasymAnd}x{\isachardot}\ P{\isacharparenleft}x{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharparenleft}Suc{\isacharparenleft}x{\isacharparenright}{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\ P{\isacharparenleft}n{\isacharparenright}{\isachardoublequote}\isanewline -\ \ Suc{\isacharunderscore}inject{\isacharcolon}\ {\isachardoublequote}Suc{\isacharparenleft}m{\isacharparenright}\ {\isacharequal}\ Suc{\isacharparenleft}n{\isacharparenright}\ {\isasymLongrightarrow}\ m\ {\isacharequal}\ n{\isachardoublequote}\isanewline -\ \ Suc{\isacharunderscore}neq{\isacharunderscore}{\isadigit{0}}{\isacharcolon}\ {\isachardoublequote}Suc{\isacharparenleft}m{\isacharparenright}\ {\isacharequal}\ {\isasymzero}\ {\isasymLongrightarrow}\ R{\isachardoublequote}\isanewline -\ \ rec{\isacharunderscore}{\isadigit{0}}{\isacharcolon}\ {\isachardoublequote}rec{\isacharparenleft}{\isasymzero}{\isacharcomma}\ a{\isacharcomma}\ f{\isacharparenright}\ {\isacharequal}\ a{\isachardoublequote}\isanewline -\ \ rec{\isacharunderscore}Suc{\isacharcolon}\ {\isachardoublequote}rec{\isacharparenleft}Suc{\isacharparenleft}m{\isacharparenright}{\isacharcomma}\ a{\isacharcomma}\ f{\isacharparenright}\ {\isacharequal}\ f{\isacharparenleft}m{\isacharcomma}\ rec{\isacharparenleft}m{\isacharcomma}\ a{\isacharcomma}\ f{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isanewline -\isanewline -\isamarkupfalse% -\isacommand{constdefs}\isanewline -\ \ add\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a{\isacharcolon}{\isacharcolon}nat\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isacharplus}{\isachardoublequote}\ {\isadigit{6}}{\isadigit{0}}{\isacharparenright}\isanewline -\ \ {\isachardoublequote}m\ {\isacharplus}\ n\ {\isasymequiv}\ rec{\isacharparenleft}m{\isacharcomma}\ n{\isacharcomma}\ {\isasymlambda}x\ y{\isachardot}\ Suc{\isacharparenleft}y{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkuptrue% -% -\begin{isamarkuptext}% -This is an abstract version of the plain \isa{Nat} theory in - FOL.\footnote{See - \url{http://isabelle.in.tum.de/library/FOL/ex/Nat.html}} Basically, - we have just replaced all occurrences of type \isa{nat} by \isa{{\isacharprime}a} and used the natural number axioms to define class \isa{nat}. - There is only a minor snag, that the original recursion operator - \isa{rec} had to be made monomorphic. - - Thus class \isa{nat} contains exactly those types \isa{{\isasymtau}} that - are isomorphic to ``the'' natural numbers (with signature \isa{{\isasymzero}}, \isa{Suc}, \isa{rec}). - - \medskip What we have done here can be also viewed as \emph{type - specification}. Of course, it still remains open if there is some - type at all that meets the class axioms. Now a very nice property of - axiomatic type classes is that abstract reasoning is always possible - --- independent of satisfiability. The meta-logic won't break, even - if some classes (or general sorts) turns out to be empty later --- - ``inconsistent'' class definitions may be useless, but do not cause - any harm. - - Theorems of the abstract natural numbers may be derived in the same - way as for the concrete version. The original proof scripts may be - re-used with some trivial changes only (mostly adding some type - constraints).% -\end{isamarkuptext}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isamarkupfalse% -\isacommand{end}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 153fe83804c9 -r 096792bdc58e doc-src/AxClass/generated/Product.tex --- a/doc-src/AxClass/generated/Product.tex Fri Aug 19 22:28:23 2005 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,132 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Product}% -\isamarkuptrue% -% -\isamarkupheader{Syntactic classes% -} -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isamarkupfalse% -\isacommand{theory}\ Product\ \isakeyword{imports}\ Main\ \isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isamarkuptrue% -% -\begin{isamarkuptext}% -\medskip\noindent There is still a feature of Isabelle's type system - left that we have not yet discussed. When declaring polymorphic - constants \isa{c\ {\isasymColon}\ {\isasymsigma}}, the type variables occurring in \isa{{\isasymsigma}} - may be constrained by type classes (or even general sorts) in an - arbitrary way. Note that by default, in Isabelle/HOL the - declaration \isa{{\isasymodot}\ {\isasymColon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a} is actually an abbreviation - for \isa{{\isasymodot}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}type\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a} Since class \isa{type} is the - universal class of HOL, this is not really a constraint at all. - - The \isa{product} class below provides a less degenerate example of - syntactic type classes.% -\end{isamarkuptext}% -\isamarkupfalse% -\isacommand{axclass}\isanewline -\ \ product\ {\isasymsubseteq}\ type\isanewline -\isamarkupfalse% -\isacommand{consts}\isanewline -\ \ product\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a{\isasymColon}product\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isasymodot}{\isachardoublequote}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isamarkuptrue% -% -\begin{isamarkuptext}% -Here class \isa{product} is defined as subclass of \isa{type} - without any additional axioms. This effects in logical equivalence - of \isa{product} and \isa{type}, as is reflected by the trivial - introduction rule generated for this definition. - - \medskip So what is the difference of declaring \isa{{\isasymodot}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}product\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a} vs.\ declaring \isa{{\isasymodot}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}type\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a} anyway? In this particular case where \isa{product\ {\isasymequiv}\ type}, it should be obvious that both declarations are the same - from the logic's point of view. It even makes the most sense to - remove sort constraints from constant declarations, as far as the - purely logical meaning is concerned \cite{Wenzel:1997:TPHOL}. - - On the other hand there are syntactic differences, of course. - Constants \isa{{\isasymodot}} on some type \isa{{\isasymtau}} are rejected by the - type-checker, unless the arity \isa{{\isasymtau}\ {\isasymColon}\ product} is part of the - type signature. In our example, this arity may be always added when - required by means of an $\INSTANCE$ with the default proof $\DDOT$. - - \medskip Thus, we may observe the following discipline of using - syntactic classes. Overloaded polymorphic constants have their type - arguments restricted to an associated (logically trivial) class - \isa{c}. Only immediately before \emph{specifying} these - constants on a certain type \isa{{\isasymtau}} do we instantiate \isa{{\isasymtau}\ {\isasymColon}\ c}. - - This is done for class \isa{product} and type \isa{bool} as - follows.% -\end{isamarkuptext}% -\isamarkupfalse% -\isacommand{instance}\ bool\ {\isacharcolon}{\isacharcolon}\ product% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isamarkupfalse% -\isacommand{{\isachardot}{\isachardot}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isamarkupfalse% -\isacommand{defs}\ {\isacharparenleft}\isakeyword{overloaded}{\isacharparenright}\isanewline -\ \ product{\isacharunderscore}bool{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequote}x\ {\isasymodot}\ y\ {\isasymequiv}\ x\ {\isasymand}\ y{\isachardoublequote}\isamarkuptrue% -% -\begin{isamarkuptext}% -The definition \isa{prod{\isacharunderscore}bool{\isacharunderscore}def} becomes syntactically - well-formed only after the arity \isa{bool\ {\isasymColon}\ product} is made - known to the type checker. - - \medskip It is very important to see that above $\DEFS$ are not - directly connected with $\INSTANCE$ at all! We were just following - our convention to specify \isa{{\isasymodot}} on \isa{bool} after having - instantiated \isa{bool\ {\isasymColon}\ product}. Isabelle does not require - these definitions, which is in contrast to programming languages like - Haskell \cite{haskell-report}. - - \medskip While Isabelle type classes and those of Haskell are almost - the same as far as type-checking and type inference are concerned, - there are important semantic differences. Haskell classes require - their instances to \emph{provide operations} of certain \emph{names}. - Therefore, its \texttt{instance} has a \texttt{where} part that tells - the system what these ``member functions'' should be. - - This style of \texttt{instance} would not make much sense in - Isabelle's meta-logic, because there is no internal notion of - ``providing operations'' or even ``names of functions''.% -\end{isamarkuptext}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isamarkupfalse% -\isacommand{end}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 153fe83804c9 -r 096792bdc58e doc-src/AxClass/generated/Semigroups.tex --- a/doc-src/AxClass/generated/Semigroups.tex Fri Aug 19 22:28:23 2005 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,88 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Semigroups}% -\isamarkuptrue% -% -\isamarkupheader{Semigroups% -} -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isamarkupfalse% -\isacommand{theory}\ Semigroups\ \isakeyword{imports}\ Main\ \isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isamarkuptrue% -% -\begin{isamarkuptext}% -\medskip\noindent An axiomatic type class is simply a class of types - that all meet certain properties, which are also called \emph{class - axioms}. Thus, type classes may be also understood as type - predicates --- i.e.\ abstractions over a single type argument \isa{{\isacharprime}a}. Class axioms typically contain polymorphic constants that - depend on this type \isa{{\isacharprime}a}. These \emph{characteristic - constants} behave like operations associated with the ``carrier'' - type \isa{{\isacharprime}a}. - - We illustrate these basic concepts by the following formulation of - semigroups.% -\end{isamarkuptext}% -\isamarkupfalse% -\isacommand{consts}\isanewline -\ \ times\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isasymodot}{\isachardoublequote}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline -\isamarkupfalse% -\isacommand{axclass}\ semigroup\ {\isasymsubseteq}\ type\isanewline -\ \ assoc{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x\ {\isasymodot}\ y{\isacharparenright}\ {\isasymodot}\ z\ {\isacharequal}\ x\ {\isasymodot}\ {\isacharparenleft}y\ {\isasymodot}\ z{\isacharparenright}{\isachardoublequote}\isamarkuptrue% -% -\begin{isamarkuptext}% -\noindent Above we have first declared a polymorphic constant \isa{{\isasymodot}\ {\isasymColon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a} and then defined the class \isa{semigroup} of - all types \isa{{\isasymtau}} such that \isa{{\isasymodot}\ {\isasymColon}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}\ {\isasymRightarrow}\ {\isasymtau}} is indeed an - associative operator. The \isa{assoc} axiom contains exactly one - type variable, which is invisible in the above presentation, though. - Also note that free term variables (like \isa{x}, \isa{y}, - \isa{z}) are allowed for user convenience --- conceptually all of - these are bound by outermost universal quantifiers. - - \medskip In general, type classes may be used to describe - \emph{structures} with exactly one carrier \isa{{\isacharprime}a} and a fixed - \emph{signature}. Different signatures require different classes. - Below, class \isa{plus{\isacharunderscore}semigroup} represents semigroups \isa{{\isacharparenleft}{\isasymtau}{\isacharcomma}\ {\isasymoplus}\isactrlsup {\isasymtau}{\isacharparenright}}, while the original \isa{semigroup} would - correspond to semigroups of the form \isa{{\isacharparenleft}{\isasymtau}{\isacharcomma}\ {\isasymodot}\isactrlsup {\isasymtau}{\isacharparenright}}.% -\end{isamarkuptext}% -\isamarkupfalse% -\isacommand{consts}\isanewline -\ \ plus\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isasymoplus}{\isachardoublequote}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline -\isamarkupfalse% -\isacommand{axclass}\ plus{\isacharunderscore}semigroup\ {\isasymsubseteq}\ type\isanewline -\ \ assoc{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x\ {\isasymoplus}\ y{\isacharparenright}\ {\isasymoplus}\ z\ {\isacharequal}\ x\ {\isasymoplus}\ {\isacharparenleft}y\ {\isasymoplus}\ z{\isacharparenright}{\isachardoublequote}\isamarkuptrue% -% -\begin{isamarkuptext}% -\noindent Even if classes \isa{plus{\isacharunderscore}semigroup} and \isa{semigroup} both represent semigroups in a sense, they are certainly - not quite the same.% -\end{isamarkuptext}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isamarkupfalse% -\isacommand{end}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r 153fe83804c9 -r 096792bdc58e doc-src/AxClass/generated/isabelle.sty --- a/doc-src/AxClass/generated/isabelle.sty Fri Aug 19 22:28:23 2005 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,191 +0,0 @@ -%% -%% Author: Markus Wenzel, TU Muenchen -%% -%% macros for Isabelle generated LaTeX output -%% -%% - -%%% Simple document preparation (based on theory token language and symbols) - -% isabelle environments - -\newcommand{\isabellecontext}{UNKNOWN} - -\newcommand{\isastyle}{\small\tt\slshape} -\newcommand{\isastyleminor}{\small\tt\slshape} -\newcommand{\isastylescript}{\footnotesize\tt\slshape} -\newcommand{\isastyletext}{\normalsize\rm} -\newcommand{\isastyletxt}{\rm} -\newcommand{\isastylecmt}{\rm} - -%symbol markup -- \emph achieves decent spacing via italic corrections -\newcommand{\isamath}[1]{\emph{$#1$}} -\newcommand{\isatext}[1]{\emph{#1}} -\newcommand{\isascriptstyle}{\def\isamath##1{##1}\def\isatext##1{\mbox{\isastylescript##1}}} -\newcommand{\isactrlsub}[1]{\emph{${}\sb{\mbox{\isastylescript{#1}}}$}} -\newcommand{\isactrlsup}[1]{\emph{${}\sp{\mbox{\isastylescript{#1}}}$}} -\newcommand{\isactrlisub}[1]{\emph{${}\sb{\mbox{\isastylescript{#1}}}$}} -\newcommand{\isactrlisup}[1]{\emph{${}\sp{\mbox{\isastylescript{#1}}}$}} -\newcommand{\isactrlbsub}{\emph\bgroup\begin{math}{}\sb\bgroup\mbox\bgroup\isastylescript} -\newcommand{\isactrlesub}{\egroup\egroup\end{math}\egroup} -\newcommand{\isactrlbsup}{\emph\bgroup\begin{math}{}\sp\bgroup\mbox\bgroup\isastylescript} -\newcommand{\isactrlesup}{\egroup\egroup\end{math}\egroup} -\newcommand{\isactrlbold}[1]{{\bfseries\upshape\boldmath#1}} - - -\newdimen\isa@parindent\newdimen\isa@parskip - -\newenvironment{isabellebody}{% -\isamarkuptrue\par% -\isa@parindent\parindent\parindent0pt% -\isa@parskip\parskip\parskip0pt% -\isastyle}{\par} - -\newenvironment{isabelle} -{\begin{trivlist}\begin{isabellebody}\item\relax} -{\end{isabellebody}\end{trivlist}} - -\newcommand{\isa}[1]{\emph{\isastyleminor #1}} - -\newcommand{\isaindent}[1]{\hphantom{#1}} -\newcommand{\isanewline}{\mbox{}\par\mbox{}} -\newcommand{\isasep}{} % override with e.g. \renewcommand{\isasep}{\vspace{1ex}} -\newcommand{\isadigit}[1]{#1} - -\chardef\isacharbang=`\! -\chardef\isachardoublequote=`\" -\chardef\isacharhash=`\# -\chardef\isachardollar=`\$ -\chardef\isacharpercent=`\% -\chardef\isacharampersand=`\& -\chardef\isacharprime=`\' -\chardef\isacharparenleft=`\( -\chardef\isacharparenright=`\) -\chardef\isacharasterisk=`\* -\chardef\isacharplus=`\+ -\chardef\isacharcomma=`\, -\chardef\isacharminus=`\- -\chardef\isachardot=`\. -\chardef\isacharslash=`\/ -\chardef\isacharcolon=`\: -\chardef\isacharsemicolon=`\; -\chardef\isacharless=`\< -\chardef\isacharequal=`\= -\chardef\isachargreater=`\> -\chardef\isacharquery=`\? -\chardef\isacharat=`\@ -\chardef\isacharbrackleft=`\[ -\chardef\isacharbackslash=`\\ -\chardef\isacharbrackright=`\] -\chardef\isacharcircum=`\^ -\chardef\isacharunderscore=`\_ -\chardef\isacharbackquote=`\` -\chardef\isacharbraceleft=`\{ -\chardef\isacharbar=`\| -\chardef\isacharbraceright=`\} -\chardef\isachartilde=`\~ - - -% keyword and section markup - -\newcommand{\isakeywordcharunderscore}{\_} -\newcommand{\isakeyword}[1] -{\emph{\bf\def\isachardot{.}\def\isacharunderscore{\isakeywordcharunderscore}% -\def\isacharbraceleft{\{}\def\isacharbraceright{\}}#1}} -\newcommand{\isacommand}[1]{\isakeyword{#1}} - -\newcommand{\isamarkupheader}[1]{\section{#1}} -\newcommand{\isamarkupchapter}[1]{\chapter{#1}} -\newcommand{\isamarkupsection}[1]{\section{#1}} -\newcommand{\isamarkupsubsection}[1]{\subsection{#1}} -\newcommand{\isamarkupsubsubsection}[1]{\subsubsection{#1}} -\newcommand{\isamarkupsect}[1]{\section{#1}} -\newcommand{\isamarkupsubsect}[1]{\subsection{#1}} -\newcommand{\isamarkupsubsubsect}[1]{\subsubsection{#1}} - -\newif\ifisamarkup -\newcommand{\isabeginpar}{\par\ifisamarkup\relax\else\medskip\fi} -\newcommand{\isaendpar}{\par\medskip} -\newenvironment{isapar}{\parindent\isa@parindent\parskip\isa@parskip\isabeginpar}{\isaendpar} -\newenvironment{isamarkuptext}{\isastyletext\begin{isapar}}{\end{isapar}} -\newenvironment{isamarkuptxt}{\isastyletxt\begin{isapar}}{\end{isapar}} -\newcommand{\isamarkupcmt}[1]{{\isastylecmt--- #1}} - - -% alternative styles - -\newcommand{\isabellestyle}{} -\def\isabellestyle#1{\csname isabellestyle#1\endcsname} - -\newcommand{\isabellestylett}{% -\renewcommand{\isastyle}{\small\tt}% -\renewcommand{\isastyleminor}{\small\tt}% -\renewcommand{\isastylescript}{\footnotesize\tt}% -} -\newcommand{\isabellestyleit}{% -\renewcommand{\isastyle}{\small\it}% -\renewcommand{\isastyleminor}{\it}% -\renewcommand{\isastylescript}{\footnotesize\it}% -\renewcommand{\isakeywordcharunderscore}{\mbox{-}}% -\renewcommand{\isacharbang}{\isamath{!}}% -\renewcommand{\isachardoublequote}{}% -\renewcommand{\isacharhash}{\isamath{\#}}% -\renewcommand{\isachardollar}{\isamath{\$}}% -\renewcommand{\isacharpercent}{\isamath{\%}}% -\renewcommand{\isacharampersand}{\isamath{\&}}% -\renewcommand{\isacharprime}{\isamath{\mskip2mu{'}\mskip-2mu}}% -\renewcommand{\isacharparenleft}{\isamath{(}}% -\renewcommand{\isacharparenright}{\isamath{)}}% -\renewcommand{\isacharasterisk}{\isamath{*}}% -\renewcommand{\isacharplus}{\isamath{+}}% -\renewcommand{\isacharcomma}{\isamath{\mathord,}}% -\renewcommand{\isacharminus}{\isamath{-}}% -\renewcommand{\isachardot}{\isamath{\mathord.}}% -\renewcommand{\isacharslash}{\isamath{/}}% -\renewcommand{\isacharcolon}{\isamath{\mathord:}}% -\renewcommand{\isacharsemicolon}{\isamath{\mathord;}}% -\renewcommand{\isacharless}{\isamath{<}}% -\renewcommand{\isacharequal}{\isamath{=}}% -\renewcommand{\isachargreater}{\isamath{>}}% -\renewcommand{\isacharat}{\isamath{@}}% -\renewcommand{\isacharbrackleft}{\isamath{[}}% -\renewcommand{\isacharbackslash}{\isamath{\backslash}}% -\renewcommand{\isacharbrackright}{\isamath{]}}% -\renewcommand{\isacharunderscore}{\mbox{-}}% -\renewcommand{\isacharbraceleft}{\isamath{\{}}% -\renewcommand{\isacharbar}{\isamath{\mid}}% -\renewcommand{\isacharbraceright}{\isamath{\}}}% -\renewcommand{\isachartilde}{\isamath{{}\sp{\sim}}}% -} - -\newcommand{\isabellestylesl}{% -\isabellestyleit% -\renewcommand{\isastyle}{\small\sl}% -\renewcommand{\isastyleminor}{\sl}% -\renewcommand{\isastylescript}{\footnotesize\sl}% -} - - -% tagged regions - -%plain TeX version of comment package -- much faster! -\let\isafmtname\fmtname\def\fmtname{plain} -\usepackage{comment} -\let\fmtname\isafmtname - -\newcommand{\isafold}[1]{\emph{$\langle\mathord{\mathit{#1}}\rangle$}} - -\newcommand{\isakeeptag}[1]% -{\includecomment{isadelim#1}\includecomment{isatag#1}\csarg\def{isafold#1}{}} -\newcommand{\isadroptag}[1]% -{\excludecomment{isadelim#1}\excludecomment{isatag#1}\csarg\def{isafold#1}{}} -\newcommand{\isafoldtag}[1]% -{\includecomment{isadelim#1}\excludecomment{isatag#1}\csarg\def{isafold#1}{\isafold{#1}}} - -\isakeeptag{theory} -\isakeeptag{proof} -\isakeeptag{ML} -\isakeeptag{visible} -\isadroptag{invisible} - -\IfFileExists{isabelletags.sty}{\usepackage{isabelletags}}{} diff -r 153fe83804c9 -r 096792bdc58e doc-src/AxClass/generated/isabellesym.sty --- a/doc-src/AxClass/generated/isabellesym.sty Fri Aug 19 22:28:23 2005 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,362 +0,0 @@ -%% -%% Author: Markus Wenzel, TU Muenchen -%% -%% definitions of standard Isabelle symbols -%% -%% - -% symbol definitions - -\newcommand{\isasymzero}{\isamath{\mathbf{0}}} %requires amssymb -\newcommand{\isasymone}{\isamath{\mathbf{1}}} %requires amssymb -\newcommand{\isasymtwo}{\isamath{\mathbf{2}}} %requires amssymb -\newcommand{\isasymthree}{\isamath{\mathbf{3}}} %requires amssymb -\newcommand{\isasymfour}{\isamath{\mathbf{4}}} %requires amssymb -\newcommand{\isasymfive}{\isamath{\mathbf{5}}} %requires amssymb -\newcommand{\isasymsix}{\isamath{\mathbf{6}}} %requires amssymb -\newcommand{\isasymseven}{\isamath{\mathbf{7}}} %requires amssymb -\newcommand{\isasymeight}{\isamath{\mathbf{8}}} %requires amssymb -\newcommand{\isasymnine}{\isamath{\mathbf{9}}} %requires amssymb -\newcommand{\isasymA}{\isamath{\mathcal{A}}} -\newcommand{\isasymB}{\isamath{\mathcal{B}}} -\newcommand{\isasymC}{\isamath{\mathcal{C}}} -\newcommand{\isasymD}{\isamath{\mathcal{D}}} -\newcommand{\isasymE}{\isamath{\mathcal{E}}} -\newcommand{\isasymF}{\isamath{\mathcal{F}}} -\newcommand{\isasymG}{\isamath{\mathcal{G}}} -\newcommand{\isasymH}{\isamath{\mathcal{H}}} -\newcommand{\isasymI}{\isamath{\mathcal{I}}} -\newcommand{\isasymJ}{\isamath{\mathcal{J}}} -\newcommand{\isasymK}{\isamath{\mathcal{K}}} -\newcommand{\isasymL}{\isamath{\mathcal{L}}} -\newcommand{\isasymM}{\isamath{\mathcal{M}}} -\newcommand{\isasymN}{\isamath{\mathcal{N}}} -\newcommand{\isasymO}{\isamath{\mathcal{O}}} -\newcommand{\isasymP}{\isamath{\mathcal{P}}} -\newcommand{\isasymQ}{\isamath{\mathcal{Q}}} -\newcommand{\isasymR}{\isamath{\mathcal{R}}} -\newcommand{\isasymS}{\isamath{\mathcal{S}}} -\newcommand{\isasymT}{\isamath{\mathcal{T}}} -\newcommand{\isasymU}{\isamath{\mathcal{U}}} -\newcommand{\isasymV}{\isamath{\mathcal{V}}} -\newcommand{\isasymW}{\isamath{\mathcal{W}}} -\newcommand{\isasymX}{\isamath{\mathcal{X}}} -\newcommand{\isasymY}{\isamath{\mathcal{Y}}} -\newcommand{\isasymZ}{\isamath{\mathcal{Z}}} -\newcommand{\isasyma}{\isamath{\mathrm{a}}} -\newcommand{\isasymb}{\isamath{\mathrm{b}}} -\newcommand{\isasymc}{\isamath{\mathrm{c}}} -\newcommand{\isasymd}{\isamath{\mathrm{d}}} -\newcommand{\isasyme}{\isamath{\mathrm{e}}} -\newcommand{\isasymf}{\isamath{\mathrm{f}}} -\newcommand{\isasymg}{\isamath{\mathrm{g}}} -\newcommand{\isasymh}{\isamath{\mathrm{h}}} -\newcommand{\isasymi}{\isamath{\mathrm{i}}} -\newcommand{\isasymj}{\isamath{\mathrm{j}}} -\newcommand{\isasymk}{\isamath{\mathrm{k}}} -\newcommand{\isasyml}{\isamath{\mathrm{l}}} -\newcommand{\isasymm}{\isamath{\mathrm{m}}} -\newcommand{\isasymn}{\isamath{\mathrm{n}}} -\newcommand{\isasymo}{\isamath{\mathrm{o}}} -\newcommand{\isasymp}{\isamath{\mathrm{p}}} -\newcommand{\isasymq}{\isamath{\mathrm{q}}} -\newcommand{\isasymr}{\isamath{\mathrm{r}}} -\newcommand{\isasyms}{\isamath{\mathrm{s}}} -\newcommand{\isasymt}{\isamath{\mathrm{t}}} -\newcommand{\isasymu}{\isamath{\mathrm{u}}} -\newcommand{\isasymv}{\isamath{\mathrm{v}}} -\newcommand{\isasymw}{\isamath{\mathrm{w}}} -\newcommand{\isasymx}{\isamath{\mathrm{x}}} -\newcommand{\isasymy}{\isamath{\mathrm{y}}} -\newcommand{\isasymz}{\isamath{\mathrm{z}}} -\newcommand{\isasymAA}{\isamath{\mathfrak{A}}} %requires eufrak -\newcommand{\isasymBB}{\isamath{\mathfrak{B}}} %requires eufrak -\newcommand{\isasymCC}{\isamath{\mathfrak{C}}} %requires eufrak -\newcommand{\isasymDD}{\isamath{\mathfrak{D}}} %requires eufrak -\newcommand{\isasymEE}{\isamath{\mathfrak{E}}} %requires eufrak -\newcommand{\isasymFF}{\isamath{\mathfrak{F}}} %requires eufrak -\newcommand{\isasymGG}{\isamath{\mathfrak{G}}} %requires eufrak -\newcommand{\isasymHH}{\isamath{\mathfrak{H}}} %requires eufrak -\newcommand{\isasymII}{\isamath{\mathfrak{I}}} %requires eufrak -\newcommand{\isasymJJ}{\isamath{\mathfrak{J}}} %requires eufrak -\newcommand{\isasymKK}{\isamath{\mathfrak{K}}} %requires eufrak -\newcommand{\isasymLL}{\isamath{\mathfrak{L}}} %requires eufrak -\newcommand{\isasymMM}{\isamath{\mathfrak{M}}} %requires eufrak -\newcommand{\isasymNN}{\isamath{\mathfrak{N}}} %requires eufrak -\newcommand{\isasymOO}{\isamath{\mathfrak{O}}} %requires eufrak -\newcommand{\isasymPP}{\isamath{\mathfrak{P}}} %requires eufrak -\newcommand{\isasymQQ}{\isamath{\mathfrak{Q}}} %requires eufrak -\newcommand{\isasymRR}{\isamath{\mathfrak{R}}} %requires eufrak -\newcommand{\isasymSS}{\isamath{\mathfrak{S}}} %requires eufrak -\newcommand{\isasymTT}{\isamath{\mathfrak{T}}} %requires eufrak -\newcommand{\isasymUU}{\isamath{\mathfrak{U}}} %requires eufrak -\newcommand{\isasymVV}{\isamath{\mathfrak{V}}} %requires eufrak -\newcommand{\isasymWW}{\isamath{\mathfrak{W}}} %requires eufrak -\newcommand{\isasymXX}{\isamath{\mathfrak{X}}} %requires eufrak -\newcommand{\isasymYY}{\isamath{\mathfrak{Y}}} %requires eufrak -\newcommand{\isasymZZ}{\isamath{\mathfrak{Z}}} %requires eufrak -\newcommand{\isasymaa}{\isamath{\mathfrak{a}}} %requires eufrak -\newcommand{\isasymbb}{\isamath{\mathfrak{b}}} %requires eufrak -\newcommand{\isasymcc}{\isamath{\mathfrak{c}}} %requires eufrak -\newcommand{\isasymdd}{\isamath{\mathfrak{d}}} %requires eufrak -\newcommand{\isasymee}{\isamath{\mathfrak{e}}} %requires eufrak -\newcommand{\isasymff}{\isamath{\mathfrak{f}}} %requires eufrak -\newcommand{\isasymgg}{\isamath{\mathfrak{g}}} %requires eufrak -\newcommand{\isasymhh}{\isamath{\mathfrak{h}}} %requires eufrak -\newcommand{\isasymii}{\isamath{\mathfrak{i}}} %requires eufrak -\newcommand{\isasymjj}{\isamath{\mathfrak{j}}} %requires eufrak -\newcommand{\isasymkk}{\isamath{\mathfrak{k}}} %requires eufrak -\newcommand{\isasymll}{\isamath{\mathfrak{l}}} %requires eufrak -\newcommand{\isasymmm}{\isamath{\mathfrak{m}}} %requires eufrak -\newcommand{\isasymnn}{\isamath{\mathfrak{n}}} %requires eufrak -\newcommand{\isasymoo}{\isamath{\mathfrak{o}}} %requires eufrak -\newcommand{\isasympp}{\isamath{\mathfrak{p}}} %requires eufrak -\newcommand{\isasymqq}{\isamath{\mathfrak{q}}} %requires eufrak -\newcommand{\isasymrr}{\isamath{\mathfrak{r}}} %requires eufrak -\newcommand{\isasymss}{\isamath{\mathfrak{s}}} %requires eufrak -\newcommand{\isasymtt}{\isamath{\mathfrak{t}}} %requires eufrak -\newcommand{\isasymuu}{\isamath{\mathfrak{u}}} %requires eufrak -\newcommand{\isasymvv}{\isamath{\mathfrak{v}}} %requires eufrak -\newcommand{\isasymww}{\isamath{\mathfrak{w}}} %requires eufrak -\newcommand{\isasymxx}{\isamath{\mathfrak{x}}} %requires eufrak -\newcommand{\isasymyy}{\isamath{\mathfrak{y}}} %requires eufrak -\newcommand{\isasymzz}{\isamath{\mathfrak{z}}} %requires eufrak -\newcommand{\isasymalpha}{\isamath{\alpha}} -\newcommand{\isasymbeta}{\isamath{\beta}} -\newcommand{\isasymgamma}{\isamath{\gamma}} -\newcommand{\isasymdelta}{\isamath{\delta}} -\newcommand{\isasymepsilon}{\isamath{\varepsilon}} -\newcommand{\isasymzeta}{\isamath{\zeta}} -\newcommand{\isasymeta}{\isamath{\eta}} -\newcommand{\isasymtheta}{\isamath{\vartheta}} -\newcommand{\isasymiota}{\isamath{\iota}} -\newcommand{\isasymkappa}{\isamath{\kappa}} -\newcommand{\isasymlambda}{\isamath{\lambda}} -\newcommand{\isasymmu}{\isamath{\mu}} -\newcommand{\isasymnu}{\isamath{\nu}} -\newcommand{\isasymxi}{\isamath{\xi}} -\newcommand{\isasympi}{\isamath{\pi}} -\newcommand{\isasymrho}{\isamath{\varrho}} -\newcommand{\isasymsigma}{\isamath{\sigma}} -\newcommand{\isasymtau}{\isamath{\tau}} -\newcommand{\isasymupsilon}{\isamath{\upsilon}} -\newcommand{\isasymphi}{\isamath{\varphi}} -\newcommand{\isasymchi}{\isamath{\chi}} -\newcommand{\isasympsi}{\isamath{\psi}} -\newcommand{\isasymomega}{\isamath{\omega}} -\newcommand{\isasymGamma}{\isamath{\Gamma}} -\newcommand{\isasymDelta}{\isamath{\Delta}} -\newcommand{\isasymTheta}{\isamath{\Theta}} -\newcommand{\isasymLambda}{\isamath{\Lambda}} -\newcommand{\isasymXi}{\isamath{\Xi}} -\newcommand{\isasymPi}{\isamath{\Pi}} -\newcommand{\isasymSigma}{\isamath{\Sigma}} -\newcommand{\isasymUpsilon}{\isamath{\Upsilon}} -\newcommand{\isasymPhi}{\isamath{\Phi}} -\newcommand{\isasymPsi}{\isamath{\Psi}} -\newcommand{\isasymOmega}{\isamath{\Omega}} -\newcommand{\isasymbool}{\isamath{\mathrm{I}\mkern-3.8mu\mathrm{B}}} -\newcommand{\isasymcomplex}{\isamath{\mathrm{C}\mkern-15mu{\phantom{\mathrm{t}}\vrule}\mkern9mu}} -\newcommand{\isasymnat}{\isamath{\mathrm{I}\mkern-3.8mu\mathrm{N}}} -\newcommand{\isasymrat}{\isamath{\mathrm{Q}\mkern-16mu{\phantom{\mathrm{t}}\vrule}\mkern10mu}} -\newcommand{\isasymreal}{\isamath{\mathrm{I}\mkern-3.8mu\mathrm{R}}} -\newcommand{\isasymint}{\isamath{\mathsf{Z}\mkern-7.5mu\mathsf{Z}}} -\newcommand{\isasymleftarrow}{\isamath{\leftarrow}} -\newcommand{\isasymlongleftarrow}{\isamath{\longleftarrow}} -\newcommand{\isasymrightarrow}{\isamath{\rightarrow}} -\newcommand{\isasymlongrightarrow}{\isamath{\longrightarrow}} -\newcommand{\isasymLeftarrow}{\isamath{\Leftarrow}} -\newcommand{\isasymLongleftarrow}{\isamath{\Longleftarrow}} -\newcommand{\isasymRightarrow}{\isamath{\Rightarrow}} -\newcommand{\isasymLongrightarrow}{\isamath{\Longrightarrow}} -\newcommand{\isasymleftrightarrow}{\isamath{\leftrightarrow}} -\newcommand{\isasymlongleftrightarrow}{\isamath{\longleftrightarrow}} -\newcommand{\isasymLeftrightarrow}{\isamath{\Leftrightarrow}} -\newcommand{\isasymLongleftrightarrow}{\isamath{\Longleftrightarrow}} -\newcommand{\isasymmapsto}{\isamath{\mapsto}} -\newcommand{\isasymlongmapsto}{\isamath{\longmapsto}} -\newcommand{\isasymmidarrow}{\isamath{\relbar}} -\newcommand{\isasymMidarrow}{\isamath{\Relbar}} -\newcommand{\isasymhookleftarrow}{\isamath{\hookleftarrow}} -\newcommand{\isasymhookrightarrow}{\isamath{\hookrightarrow}} -\newcommand{\isasymleftharpoondown}{\isamath{\leftharpoondown}} -\newcommand{\isasymrightharpoondown}{\isamath{\rightharpoondown}} -\newcommand{\isasymleftharpoonup}{\isamath{\leftharpoonup}} -\newcommand{\isasymrightharpoonup}{\isamath{\rightharpoonup}} -\newcommand{\isasymrightleftharpoons}{\isamath{\rightleftharpoons}} -\newcommand{\isasymdownharpoonleft}{\isamath{\downharpoonleft}} %requires amssymb -\newcommand{\isasymdownharpoonright}{\isamath{\downharpoonright}} %requires amssymb -\newcommand{\isasymupharpoonleft}{\isamath{\upharpoonleft}} %requires amssymb -\newcommand{\isasymupharpoonright}{\isamath{\upharpoonright}} %requires amssymb -\newcommand{\isasymrestriction}{\isamath{\restriction}} %requires amssymb -\newcommand{\isasymleadsto}{\isamath{\leadsto}} %requires amssymb -\newcommand{\isasymup}{\isamath{\uparrow}} -\newcommand{\isasymUp}{\isamath{\Uparrow}} -\newcommand{\isasymdown}{\isamath{\downarrow}} -\newcommand{\isasymDown}{\isamath{\Downarrow}} -\newcommand{\isasymupdown}{\isamath{\updownarrow}} -\newcommand{\isasymUpdown}{\isamath{\Updownarrow}} -\newcommand{\isasymlangle}{\isamath{\langle}} -\newcommand{\isasymrangle}{\isamath{\rangle}} -\newcommand{\isasymlceil}{\isamath{\lceil}} -\newcommand{\isasymrceil}{\isamath{\rceil}} -\newcommand{\isasymlfloor}{\isamath{\lfloor}} -\newcommand{\isasymrfloor}{\isamath{\rfloor}} -\newcommand{\isasymlparr}{\isamath{\mathopen{(\mkern-3mu\mid}}} -\newcommand{\isasymrparr}{\isamath{\mathclose{\mid\mkern-3mu)}}} -\newcommand{\isasymlbrakk}{\isamath{\mathopen{\lbrack\mkern-3mu\lbrack}}} -\newcommand{\isasymrbrakk}{\isamath{\mathclose{\rbrack\mkern-3mu\rbrack}}} -\newcommand{\isasymlbrace}{\isamath{\mathopen{\lbrace\mkern-4.5mu\mid}}} -\newcommand{\isasymrbrace}{\isamath{\mathclose{\mid\mkern-4.5mu\rbrace}}} -\newcommand{\isasymguillemotleft}{\isatext{\flqq}} %requires babel -\newcommand{\isasymguillemotright}{\isatext{\frqq}} %requires babel -\newcommand{\isasymColon}{\isamath{\mathrel{::}}} -\newcommand{\isasymnot}{\isamath{\neg}} -\newcommand{\isasymbottom}{\isamath{\bot}} -\newcommand{\isasymtop}{\isamath{\top}} -\newcommand{\isasymand}{\isamath{\wedge}} -\newcommand{\isasymAnd}{\isamath{\bigwedge}} -\newcommand{\isasymor}{\isamath{\vee}} -\newcommand{\isasymOr}{\isamath{\bigvee}} -\newcommand{\isasymforall}{\isamath{\forall\,}} -\newcommand{\isasymexists}{\isamath{\exists\,}} -\newcommand{\isasymnexists}{\isamath{\nexists\,}} %requires amssymb -\newcommand{\isasymbox}{\isamath{\Box}} %requires amssymb -\newcommand{\isasymdiamond}{\isamath{\Diamond}} %requires amssymb -\newcommand{\isasymturnstile}{\isamath{\vdash}} -\newcommand{\isasymTurnstile}{\isamath{\models}} -\newcommand{\isasymtturnstile}{\isamath{\vdash\!\!\!\vdash}} -\newcommand{\isasymTTurnstile}{\isamath{\mid\!\models}} -\newcommand{\isasymstileturn}{\isamath{\dashv}} -\newcommand{\isasymsurd}{\isamath{\surd}} -\newcommand{\isasymle}{\isamath{\le}} -\newcommand{\isasymge}{\isamath{\ge}} -\newcommand{\isasymlless}{\isamath{\ll}} -\newcommand{\isasymggreater}{\isamath{\gg}} -\newcommand{\isasymlesssim}{\isamath{\lesssim}} %requires amssymb -\newcommand{\isasymgreatersim}{\isamath{\gtrsim}} %requires amssymb -\newcommand{\isasymlessapprox}{\isamath{\lessapprox}} %requires amssymb -\newcommand{\isasymgreaterapprox}{\isamath{\gtrapprox}} %requires amssymb -\newcommand{\isasymin}{\isamath{\in}} -\newcommand{\isasymnotin}{\isamath{\notin}} -\newcommand{\isasymsubset}{\isamath{\subset}} -\newcommand{\isasymsupset}{\isamath{\supset}} -\newcommand{\isasymsubseteq}{\isamath{\subseteq}} -\newcommand{\isasymsupseteq}{\isamath{\supseteq}} -\newcommand{\isasymsqsubset}{\isamath{\sqsubset}} %requires amssymb -\newcommand{\isasymsqsupset}{\isamath{\sqsupset}} %requires amssymb -\newcommand{\isasymsqsubseteq}{\isamath{\sqsubseteq}} -\newcommand{\isasymsqsupseteq}{\isamath{\sqsupseteq}} -\newcommand{\isasyminter}{\isamath{\cap}} -\newcommand{\isasymInter}{\isamath{\bigcap\,}} -\newcommand{\isasymunion}{\isamath{\cup}} -\newcommand{\isasymUnion}{\isamath{\bigcup\,}} -\newcommand{\isasymsqunion}{\isamath{\sqcup}} -\newcommand{\isasymSqunion}{\isamath{\bigsqcup\,}} -\newcommand{\isasymsqinter}{\isamath{\sqcap}} -\newcommand{\isasymSqinter}{\isamath{\bigsqcap\,}} %requires amsmath -\newcommand{\isasymuplus}{\isamath{\uplus}} -\newcommand{\isasymUplus}{\isamath{\biguplus\,}} -\newcommand{\isasymnoteq}{\isamath{\not=}} -\newcommand{\isasymsim}{\isamath{\sim}} -\newcommand{\isasymdoteq}{\isamath{\doteq}} -\newcommand{\isasymsimeq}{\isamath{\simeq}} -\newcommand{\isasymapprox}{\isamath{\approx}} -\newcommand{\isasymasymp}{\isamath{\asymp}} -\newcommand{\isasymcong}{\isamath{\cong}} -\newcommand{\isasymsmile}{\isamath{\smile}} -\newcommand{\isasymequiv}{\isamath{\equiv}} -\newcommand{\isasymfrown}{\isamath{\frown}} -\newcommand{\isasympropto}{\isamath{\propto}} -\newcommand{\isasymsome}{\isamath{\epsilon\,}} -\newcommand{\isasymJoin}{\isamath{\Join}} %requires amssymb -\newcommand{\isasymbowtie}{\isamath{\bowtie}} -\newcommand{\isasymprec}{\isamath{\prec}} -\newcommand{\isasymsucc}{\isamath{\succ}} -\newcommand{\isasympreceq}{\isamath{\preceq}} -\newcommand{\isasymsucceq}{\isamath{\succeq}} -\newcommand{\isasymparallel}{\isamath{\parallel}} -\newcommand{\isasymbar}{\isamath{\mid}} -\newcommand{\isasymplusminus}{\isamath{\pm}} -\newcommand{\isasymminusplus}{\isamath{\mp}} -\newcommand{\isasymtimes}{\isamath{\times}} -\newcommand{\isasymdiv}{\isamath{\div}} -\newcommand{\isasymcdot}{\isamath{\cdot}} -\newcommand{\isasymstar}{\isamath{\star}} -\newcommand{\isasymbullet}{\boldmath\isamath{\mathchoice{\displaystyle{\cdot}}{\textstyle{\cdot}}{\scriptstyle{\bullet}}{\scriptscriptstyle{\bullet}}}} -\newcommand{\isasymcirc}{\isamath{\circ}} -\newcommand{\isasymdagger}{\isamath{\dagger}} -\newcommand{\isasymddagger}{\isamath{\ddagger}} -\newcommand{\isasymlhd}{\isamath{\lhd}} %requires amssymb -\newcommand{\isasymrhd}{\isamath{\rhd}} %requires amssymb -\newcommand{\isasymunlhd}{\isamath{\unlhd}} %requires amssymb -\newcommand{\isasymunrhd}{\isamath{\unrhd}} %requires amssymb -\newcommand{\isasymtriangleleft}{\isamath{\triangleleft}} -\newcommand{\isasymtriangleright}{\isamath{\triangleright}} -\newcommand{\isasymtriangle}{\isamath{\triangle}} -\newcommand{\isasymtriangleq}{\isamath{\triangleq}} %requires amssymb -\newcommand{\isasymoplus}{\isamath{\oplus}} -\newcommand{\isasymOplus}{\isamath{\bigoplus\,}} -\newcommand{\isasymotimes}{\isamath{\otimes}} -\newcommand{\isasymOtimes}{\isamath{\bigotimes\,}} -\newcommand{\isasymodot}{\isamath{\odot}} -\newcommand{\isasymOdot}{\isamath{\bigodot\,}} -\newcommand{\isasymominus}{\isamath{\ominus}} -\newcommand{\isasymoslash}{\isamath{\oslash}} -\newcommand{\isasymdots}{\isamath{\dots}} -\newcommand{\isasymcdots}{\isamath{\cdots}} -\newcommand{\isasymSum}{\isamath{\sum\,}} -\newcommand{\isasymProd}{\isamath{\prod\,}} -\newcommand{\isasymCoprod}{\isamath{\coprod\,}} -\newcommand{\isasyminfinity}{\isamath{\infty}} -\newcommand{\isasymintegral}{\isamath{\int\,}} -\newcommand{\isasymointegral}{\isamath{\oint\,}} -\newcommand{\isasymclubsuit}{\isamath{\clubsuit}} -\newcommand{\isasymdiamondsuit}{\isamath{\diamondsuit}} -\newcommand{\isasymheartsuit}{\isamath{\heartsuit}} -\newcommand{\isasymspadesuit}{\isamath{\spadesuit}} -\newcommand{\isasymaleph}{\isamath{\aleph}} -\newcommand{\isasymemptyset}{\isamath{\emptyset}} -\newcommand{\isasymnabla}{\isamath{\nabla}} -\newcommand{\isasympartial}{\isamath{\partial}} -\newcommand{\isasymRe}{\isamath{\Re}} -\newcommand{\isasymIm}{\isamath{\Im}} -\newcommand{\isasymflat}{\isamath{\flat}} -\newcommand{\isasymnatural}{\isamath{\natural}} -\newcommand{\isasymsharp}{\isamath{\sharp}} -\newcommand{\isasymangle}{\isamath{\angle}} -\newcommand{\isasymcopyright}{\isatext{\rm\copyright}} -\newcommand{\isasymregistered}{\isatext{\rm\textregistered}} -\newcommand{\isasymhyphen}{\isatext{\rm-}} -\newcommand{\isasyminverse}{\isamath{{}^{-1}}} -\newcommand{\isasymonesuperior}{\isamath{\mathonesuperior}} %requires latin1 -\newcommand{\isasymonequarter}{\isatext{\rm\textonequarter}} %requires latin1 -\newcommand{\isasymtwosuperior}{\isamath{\mathtwosuperior}} %requires latin1 -\newcommand{\isasymonehalf}{\isatext{\rm\textonehalf}} %requires latin1 -\newcommand{\isasymthreesuperior}{\isamath{\maththreesuperior}} %requires latin1 -\newcommand{\isasymthreequarters}{\isatext{\rm\textthreequarters}} %requires latin1 -\newcommand{\isasymordfeminine}{\isatext{\rm\textordfeminine}} -\newcommand{\isasymordmasculine}{\isatext{\rm\textordmasculine}} -\newcommand{\isasymsection}{\isatext{\rm\S}} -\newcommand{\isasymparagraph}{\isatext{\rm\P}} -\newcommand{\isasymexclamdown}{\isatext{\rm\textexclamdown}} -\newcommand{\isasymquestiondown}{\isatext{\rm\textquestiondown}} -\newcommand{\isasymeuro}{\isatext{\textgreek{\euro}}} %requires greek babel -\newcommand{\isasympounds}{\isamath{\pounds}} -\newcommand{\isasymyen}{\isatext{\yen}} %requires amssymb -\newcommand{\isasymcent}{\isatext{\textcent}} %requires textcomp -\newcommand{\isasymcurrency}{\isatext{\textcurrency}} %requires textcomp -\newcommand{\isasymdegree}{\isatext{\rm\textdegree}} %requires latin1 -\newcommand{\isasymamalg}{\isamath{\amalg}} -\newcommand{\isasymmho}{\isamath{\mho}} %requires amssymb -\newcommand{\isasymlozenge}{\isamath{\lozenge}} %requires amssymb -\newcommand{\isasymwp}{\isamath{\wp}} -\newcommand{\isasymwrong}{\isamath{\wr}} -\newcommand{\isasymstruct}{\isamath{\diamond}} -\newcommand{\isasymacute}{\isatext{\'\relax}} -\newcommand{\isasymindex}{\isatext{\i}} -\newcommand{\isasymdieresis}{\isatext{\"\relax}} -\newcommand{\isasymcedilla}{\isatext{\c\relax}} -\newcommand{\isasymhungarumlaut}{\isatext{\H\relax}} -\newcommand{\isasymspacespace}{\isamath{~~}} -\newcommand{\isasymmodule}{\isamath{\langle}\isakeyword{module}\isamath{\rangle}}