# HG changeset patch # User haftmann # Date 1160549867 -7200 # Node ID 1df105407f8781edfa25a74f72d29ecb9d863b0f # Parent 75c8a52f84475fd4178d8639e60e4e31a28de166 added tex files to CVS diff -r 75c8a52f8447 -r 1df105407f87 doc-src/IsarAdvanced/Classes/Thy/document/Classes.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarAdvanced/Classes/Thy/document/Classes.tex Wed Oct 11 08:57:47 2006 +0200 @@ -0,0 +1,688 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Classes}% +% +\isadelimtheory +\isanewline +\isanewline +\isanewline +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Classes\isanewline +\isakeyword{imports}\ Main\isanewline +\isakeyword{begin}\isanewline +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isadelimML +% +\endisadelimML +% +\isatagML +% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\isamarkupchapter{Haskell-style classes with Isabelle/Isar% +} +\isamarkuptrue% +% +\isamarkupsection{Introduction% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The well-known concept of type classes + \cite{wadler89how,peterson93implementing,hall96type,Nipkow-Prehofer:1993,Nipkow:1993,Wenzel:1997} + offers a useful structuring mechanism for programs and proofs, which + is more light-weight than a fully featured module mechanism. Type + classes are able to qualify types by associating operations and + logical properties. For example, class \isa{eq} could provide + an equivalence relation \isa{{\isacharequal}} on type \isa{{\isasymalpha}}, and class + \isa{ord} could extend \isa{eq} by providing a strict order + \isa{{\isacharless}} etc. + + Isabelle/Isar offers Haskell-style type classes, combining operational + and logical specifications.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{A simple algebra example \label{sec:example}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +We demonstrate common elements of structured specifications and + abstract reasoning with type classes by the algebraic hierarchy of + semigroups, monoids and groups. Our background theory is that of + Isabelle/HOL \cite{Nipkow-et-al:2002:tutorial}, which uses fairly + standard notation from mathematics and functional programming. We + also refer to basic vernacular commands for definitions and + statements, e.g.\ \isa{{\isasymDEFINITION}} and \isa{{\isasymLEMMA}}; + proofs will be recorded using structured elements of Isabelle/Isar + \cite{Wenzel-PhD,Nipkow:2002}, notably \isa{{\isasymPROOF}}/\isa{{\isasymQED}} and \isa{{\isasymFIX}}/\isa{{\isasymASSUME}}/\isa{{\isasymSHOW}}. + + Our main concern are the new \isa{{\isasymCLASS}} + and \isa{{\isasymINSTANCE}} elements used below. + Here we merely present the + look-and-feel for end users, which is quite similar to Haskell's + \texttt{class} and \texttt{instance} \cite{hall96type}, but + augmented by logical specifications and proofs; + Internally, those are mapped to more primitive Isabelle concepts. + See \cite{haftmann_wenzel2006classes} for more detail.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Class definition% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Depending on an arbitrary type \isa{{\isasymalpha}}, class \isa{semigroup} introduces a binary operation \isa{{\isasymcirc}} that is + assumed to be associative:% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \ \ \isacommand{class}\isamarkupfalse% +\ semigroup\ {\isacharequal}\isanewline +\ \ \ \ \ \ \isakeyword{fixes}\ mult\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}\isactrlloc {\isasymotimes}{\isachardoublequoteclose}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline +\ \ \ \ \ \ \isakeyword{assumes}\ assoc{\isacharcolon}\ {\isachardoublequoteopen}{\isacharparenleft}x\ \isactrlloc {\isasymotimes}\ y{\isacharparenright}\ \isactrlloc {\isasymotimes}\ z\ {\isacharequal}\ x\ \isactrlloc {\isasymotimes}\ {\isacharparenleft}y\ \isactrlloc {\isasymotimes}\ z{\isacharparenright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent This \isa{{\isasymCLASS}} specification consists of two + parts: the \qn{operational} part names the class operation (\isa{{\isasymFIXES}}), the \qn{logical} part specifies properties on them + (\isa{{\isasymASSUMES}}). The local \isa{{\isasymFIXES}} and \isa{{\isasymASSUMES}} are lifted to the theory toplevel, yielding the global + operation \isa{{\isachardoublequote}mult\ {\isacharcolon}{\isacharcolon}\ {\isasymalpha}{\isacharcolon}{\isacharcolon}semigroup\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequote}} and the + global theorem \isa{semigroup{\isachardot}assoc{\isacharcolon}}~\isa{{\isachardoublequote}{\isasymAnd}x\ y\ z{\isacharcolon}{\isacharcolon}{\isasymalpha}{\isacharcolon}{\isacharcolon}semigroup{\isachardot}\ {\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isasymotimes}\ z\ {\isacharequal}\ x\ {\isasymotimes}\ {\isacharparenleft}y\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequote}}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Class instantiation \label{sec:class_inst}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +The concrete type \isa{int} is made a \isa{semigroup} + instance by providing a suitable definition for the class operation + \isa{mult} and a proof for the specification of \isa{assoc}.% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \ \ \isacommand{instance}\isamarkupfalse% +\ int\ {\isacharcolon}{\isacharcolon}\ semigroup\isanewline +\ \ \ \ \ \ \ \ mult{\isacharunderscore}int{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequoteopen}{\isasymAnd}i\ j\ {\isacharcolon}{\isacharcolon}\ int{\isachardot}\ i\ {\isasymotimes}\ j\ {\isasymequiv}\ i\ {\isacharplus}\ j{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% +\ i\ j\ k\ {\isacharcolon}{\isacharcolon}\ int\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharparenleft}i\ {\isacharplus}\ j{\isacharparenright}\ {\isacharplus}\ k\ {\isacharequal}\ i\ {\isacharplus}\ {\isacharparenleft}j\ {\isacharplus}\ k{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharparenleft}i\ {\isasymotimes}\ j{\isacharparenright}\ {\isasymotimes}\ k\ {\isacharequal}\ i\ {\isasymotimes}\ {\isacharparenleft}j\ {\isasymotimes}\ k{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{unfolding}\isamarkupfalse% +\ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent From now on, the type-checker will consider \isa{int} + as a \isa{semigroup} automatically, i.e.\ any general results + are immediately available on concrete instances. + + Another instance of \isa{semigroup} are the natural numbers:% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \ \ \isacommand{instance}\isamarkupfalse% +\ nat\ {\isacharcolon}{\isacharcolon}\ semigroup\isanewline +\ \ \ \ \ \ {\isachardoublequoteopen}m\ {\isasymotimes}\ n\ {\isasymequiv}\ m\ {\isacharplus}\ n{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% +\ m\ n\ q\ {\isacharcolon}{\isacharcolon}\ nat\ \isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}m\ {\isasymotimes}\ n\ {\isasymotimes}\ q\ {\isacharequal}\ m\ {\isasymotimes}\ {\isacharparenleft}n\ {\isasymotimes}\ q{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{unfolding}\isamarkupfalse% +\ semigroup{\isacharunderscore}nat{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +Also \isa{list}s form a semigroup with \isa{op\ {\isacharat}} as + operation:% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \ \ \isacommand{instance}\isamarkupfalse% +\ list\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}type{\isacharparenright}\ semigroup\isanewline +\ \ \ \ \ \ {\isachardoublequoteopen}xs\ {\isasymotimes}\ ys\ {\isasymequiv}\ xs\ {\isacharat}\ ys{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% +\ xs\ ys\ zs\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ list{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}xs\ {\isasymotimes}\ ys\ {\isasymotimes}\ zs\ {\isacharequal}\ xs\ {\isasymotimes}\ {\isacharparenleft}ys\ {\isasymotimes}\ zs{\isacharparenright}{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% +\ semigroup{\isacharunderscore}list{\isacharunderscore}def\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymAnd}xs\ ys{\isasymColon}{\isasymalpha}\ list{\isachardot}\ xs\ {\isasymotimes}\ ys\ {\isasymequiv}\ xs\ {\isacharat}\ ys{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{thus}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{Subclasses% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +We define a subclass \isa{monoidl} (a semigroup with an left-hand neutral) + by extending \isa{semigroup} + with one additional operation \isa{neutral} together + with its property:% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \ \ \isacommand{class}\isamarkupfalse% +\ monoidl\ {\isacharequal}\ semigroup\ {\isacharplus}\isanewline +\ \ \ \ \ \ \isakeyword{fixes}\ neutral\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}{\isachardoublequoteclose}\ {\isacharparenleft}{\isachardoublequoteopen}\isactrlloc {\isasymone}{\isachardoublequoteclose}{\isacharparenright}\isanewline +\ \ \ \ \ \ \isakeyword{assumes}\ neutl{\isacharcolon}\ {\isachardoublequoteopen}\isactrlloc {\isasymone}\ \isactrlloc {\isasymotimes}\ x\ {\isacharequal}\ x{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent Again, we make some instances, by + providing suitable operation definitions and proofs for the + additional specifications.% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \ \ \isacommand{instance}\isamarkupfalse% +\ nat\ {\isacharcolon}{\isacharcolon}\ monoidl\isanewline +\ \ \ \ \ \ {\isachardoublequoteopen}{\isasymone}\ {\isasymequiv}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% +\ n\ {\isacharcolon}{\isacharcolon}\ nat\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ n\ {\isacharequal}\ n{\isachardoublequoteclose}\ \isacommand{unfolding}\isamarkupfalse% +\ neutral{\isacharunderscore}nat{\isacharunderscore}def\ mult{\isacharunderscore}nat{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\ \ \ \ \isacommand{instance}\isamarkupfalse% +\ int\ {\isacharcolon}{\isacharcolon}\ monoidl\isanewline +\ \ \ \ \ \ {\isachardoublequoteopen}{\isasymone}\ {\isasymequiv}\ {\isadigit{0}}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% +\ k\ {\isacharcolon}{\isacharcolon}\ int\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ k\ {\isacharequal}\ k{\isachardoublequoteclose}\ \isacommand{unfolding}\isamarkupfalse% +\ neutral{\isacharunderscore}int{\isacharunderscore}def\ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\ \ \ \ \isacommand{instance}\isamarkupfalse% +\ list\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}type{\isacharparenright}\ monoidl\isanewline +\ \ \ \ \ \ {\isachardoublequoteopen}{\isasymone}\ {\isasymequiv}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% +\ xs\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ list{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymone}\ {\isasymotimes}\ xs\ {\isacharequal}\ xs{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% +\ mult{\isacharunderscore}list{\isacharunderscore}def\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymAnd}xs\ ys{\isasymColon}{\isacharprime}a\ list{\isachardot}\ xs\ {\isasymotimes}\ ys\ {\isasymequiv}\ xs\ {\isacharat}\ ys{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse% +\ \isacommand{from}\isamarkupfalse% +\ mult{\isacharunderscore}list{\isacharunderscore}def\ neutral{\isacharunderscore}list{\isacharunderscore}def\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymone}\ {\isasymequiv}\ {\isacharbrackleft}{\isacharbrackright}{\isasymColon}{\isasymalpha}\ list{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \ \ \ \ \isacommand{ultimately}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +To finish our small algebra example, we add \isa{monoid} + and \isa{group} classes with corresponding instances% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \ \ \isacommand{class}\isamarkupfalse% +\ monoid\ {\isacharequal}\ monoidl\ {\isacharplus}\isanewline +\ \ \ \ \ \ \isakeyword{assumes}\ neutr{\isacharcolon}\ {\isachardoublequoteopen}x\ \isactrlloc {\isasymotimes}\ \isactrlloc {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline +\isanewline +\ \ \ \ \isacommand{instance}\isamarkupfalse% +\ nat\ {\isacharcolon}{\isacharcolon}\ monoid\isanewline +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% +\ n\ {\isacharcolon}{\isacharcolon}\ nat\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}n\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ n{\isachardoublequoteclose}\ \isacommand{unfolding}\isamarkupfalse% +\ neutral{\isacharunderscore}nat{\isacharunderscore}def\ mult{\isacharunderscore}nat{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\ \ \ \ \isacommand{instance}\isamarkupfalse% +\ int\ {\isacharcolon}{\isacharcolon}\ monoid\isanewline +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% +\ k\ {\isacharcolon}{\isacharcolon}\ int\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}k\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ k{\isachardoublequoteclose}\ \isacommand{unfolding}\isamarkupfalse% +\ neutral{\isacharunderscore}int{\isacharunderscore}def\ mult{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\isanewline +\ \ \ \ \isacommand{instance}\isamarkupfalse% +\ list\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}type{\isacharparenright}\ monoid\isanewline +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% +\ xs\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ list{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}xs\ {\isasymotimes}\ {\isasymone}\ {\isacharequal}\ xs{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \ \ \ \ \isacommand{from}\isamarkupfalse% +\ mult{\isacharunderscore}list{\isacharunderscore}def\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymAnd}xs\ ys{\isasymColon}{\isasymalpha}\ list{\isachardot}\ xs\ {\isasymotimes}\ ys\ {\isasymequiv}\ xs\ {\isacharat}\ ys{\isachardoublequoteclose}\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \ \ \isacommand{moreover}\isamarkupfalse% +\ \isacommand{from}\isamarkupfalse% +\ mult{\isacharunderscore}list{\isacharunderscore}def\ neutral{\isacharunderscore}list{\isacharunderscore}def\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isasymone}\ {\isasymequiv}\ {\isacharbrackleft}{\isacharbrackright}{\isasymColon}{\isacharprime}a\ list{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \ \ \ \ \isacommand{ultimately}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isacharquery}thesis\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\ \ \isanewline +% +\endisadelimproof +\isanewline +\ \ \ \ \isacommand{class}\isamarkupfalse% +\ group\ {\isacharequal}\ monoidl\ {\isacharplus}\isanewline +\ \ \ \ \ \ \isakeyword{fixes}\ inverse\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}{\isachardoublequoteopen}{\isacharparenleft}{\isacharunderscore}\isactrlloc {\isasymdiv}{\isacharparenright}{\isachardoublequoteclose}\ {\isacharbrackleft}{\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isadigit{0}}{\isacharbrackright}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharparenright}\isanewline +\ \ \ \ \ \ \isakeyword{assumes}\ invl{\isacharcolon}\ {\isachardoublequoteopen}x\isactrlloc {\isasymdiv}\ \isactrlloc {\isasymotimes}\ x\ {\isacharequal}\ \isactrlloc {\isasymone}{\isachardoublequoteclose}\isanewline +\isanewline +\ \ \ \ \isacommand{instance}\isamarkupfalse% +\ int\ {\isacharcolon}{\isacharcolon}\ group\isanewline +\ \ \ \ \ \ {\isachardoublequoteopen}i{\isasymdiv}\ {\isasymequiv}\ {\isacharminus}\ i{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% +\ i\ {\isacharcolon}{\isacharcolon}\ int\isanewline +\ \ \ \ \ \ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharminus}i\ {\isacharplus}\ i\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}i{\isasymdiv}\ {\isasymotimes}\ i\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\ \isacommand{unfolding}\isamarkupfalse% +\ mult{\isacharunderscore}int{\isacharunderscore}def\ \isakeyword{and}\ neutral{\isacharunderscore}int{\isacharunderscore}def\ \isakeyword{and}\ inverse{\isacharunderscore}int{\isacharunderscore}def\ \isacommand{{\isachardot}}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsubsection{Abstract reasoning% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Abstract theories enable reasoning at a general level, while results + are implicitly transferred to all instances. For example, we can + now establish the \isa{left{\isacharunderscore}cancel} lemma for groups, which + states that the function \isa{{\isacharparenleft}x\ {\isasymcirc}{\isacharparenright}} is injective:% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \ \ \isacommand{lemma}\isamarkupfalse% +\ {\isacharparenleft}\isakeyword{in}\ group{\isacharparenright}\ left{\isacharunderscore}cancel{\isacharcolon}\ {\isachardoublequoteopen}x\ \isactrlloc {\isasymotimes}\ y\ {\isacharequal}\ x\ \isactrlloc {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ \isactrlloc {\isasymotimes}\ y\ {\isacharequal}\ x\ \isactrlloc {\isasymotimes}\ z{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\isactrlloc {\isasymdiv}\ \isactrlloc {\isasymotimes}\ {\isacharparenleft}x\ \isactrlloc {\isasymotimes}\ y{\isacharparenright}\ {\isacharequal}\ x\isactrlloc {\isasymdiv}\ \isactrlloc {\isasymotimes}\ {\isacharparenleft}x\ \isactrlloc {\isasymotimes}\ z{\isacharparenright}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}{\isacharparenleft}x\isactrlloc {\isasymdiv}\ \isactrlloc {\isasymotimes}\ x{\isacharparenright}\ \isactrlloc {\isasymotimes}\ y\ {\isacharequal}\ {\isacharparenleft}x\isactrlloc {\isasymdiv}\ \isactrlloc {\isasymotimes}\ x{\isacharparenright}\ \isactrlloc {\isasymotimes}\ z{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse% +\ assoc\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isacharequal}\ z{\isachardoublequoteclose}\ \isacommand{using}\isamarkupfalse% +\ neutl\ \isakeyword{and}\ invl\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \isacommand{next}\isamarkupfalse% +\isanewline +\ \ \ \ \isacommand{assume}\isamarkupfalse% +\ {\isachardoublequoteopen}y\ {\isacharequal}\ z{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ \ \ \isacommand{then}\isamarkupfalse% +\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ \isactrlloc {\isasymotimes}\ y\ {\isacharequal}\ x\ \isactrlloc {\isasymotimes}\ z{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\begin{isamarkuptext}% +\noindent Here the \qt{\isa{{\isasymIN}\ group}} target specification + indicates that the result is recorded within that context for later + use. This local theorem is also lifted to the global one \isa{group{\isachardot}left{\isacharunderscore}cancel{\isacharcolon}} \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z{\isacharcolon}{\isacharcolon}{\isasymalpha}{\isacharcolon}{\isacharcolon}group{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}. Since type \isa{int} has been made an instance of + \isa{group} before, we may refer to that fact as well: \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z{\isacharcolon}{\isacharcolon}int{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Additional subclass relations% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Any \isa{group} is also a \isa{monoid}; this + can be made explicit by claiming an additional subclass relation, + together with a proof of the logical difference:% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \ \ \isacommand{instance}\isamarkupfalse% +\ group\ {\isacharless}\ monoid\isanewline +% +\isadelimproof +\ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{proof}\isamarkupfalse% +\ {\isacharminus}\isanewline +\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% +\ x\isanewline +\ \ \ \ \ \ \isacommand{from}\isamarkupfalse% +\ invl\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\isactrlloc {\isasymdiv}\ \isactrlloc {\isasymotimes}\ x\ {\isacharequal}\ \isactrlloc {\isasymone}{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \ \ \isacommand{with}\isamarkupfalse% +\ assoc\ {\isacharbrackleft}symmetric{\isacharbrackright}\ neutl\ invl\ \isacommand{have}\isamarkupfalse% +\ {\isachardoublequoteopen}x\isactrlloc {\isasymdiv}\ \isactrlloc {\isasymotimes}\ {\isacharparenleft}x\ \isactrlloc {\isasymotimes}\ \isactrlloc {\isasymone}{\isacharparenright}\ {\isacharequal}\ x\isactrlloc {\isasymdiv}\ \isactrlloc {\isasymotimes}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \ \ \isacommand{with}\isamarkupfalse% +\ left{\isacharunderscore}cancel\ \isacommand{show}\isamarkupfalse% +\ {\isachardoublequoteopen}x\ \isactrlloc {\isasymotimes}\ \isactrlloc {\isasymone}\ {\isacharequal}\ x{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse% +\ simp\isanewline +\ \ \ \ \isacommand{qed}\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isamarkupsection{Code generation% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Code generation takes account of type classes, + resulting either in Haskell type classes or SML dictionaries. + As example, we define the natural power function on groups:% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \ \ \isacommand{function}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ pow{\isacharunderscore}nat\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ {\isasymRightarrow}\ {\isacharprime}a{\isasymColon}monoidl\ {\isasymRightarrow}\ {\isacharprime}a{\isasymColon}monoidl{\isachardoublequoteclose}\ \isakeyword{where}\isanewline +\ \ \ \ \ \ {\isachardoublequoteopen}pow{\isacharunderscore}nat\ {\isadigit{0}}\ x\ {\isacharequal}\ {\isasymone}{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ {\isachardoublequoteopen}pow{\isacharunderscore}nat\ {\isacharparenleft}Suc\ n{\isacharparenright}\ x\ {\isacharequal}\ x\ {\isasymotimes}\ pow{\isacharunderscore}nat\ n\ x{\isachardoublequoteclose}\isanewline +% +\isadelimproof +\ \ \ \ \ \ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ pat{\isacharunderscore}completeness\ auto% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +\ \ \ \ \isacommand{termination}\isamarkupfalse% +\ pow{\isacharunderscore}nat% +\isadelimproof +\ % +\endisadelimproof +% +\isatagproof +\isacommand{by}\isamarkupfalse% +\ {\isacharparenleft}auto{\isacharunderscore}term\ {\isachardoublequoteopen}measure\ fst{\isachardoublequoteclose}{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline +\ \ \ \ \isacommand{declare}\isamarkupfalse% +\ pow{\isacharunderscore}nat{\isachardot}simps\ {\isacharbrackleft}code\ func{\isacharbrackright}\isanewline +\isanewline +\ \ \ \ \isacommand{definition}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ pow{\isacharunderscore}int\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}int\ {\isasymRightarrow}\ {\isacharprime}a{\isasymColon}group\ {\isasymRightarrow}\ {\isacharprime}a{\isasymColon}group{\isachardoublequoteclose}\isanewline +\ \ \ \ \ \ {\isachardoublequoteopen}pow{\isacharunderscore}int\ k\ x\ {\isacharequal}\ {\isacharparenleft}if\ k\ {\isachargreater}{\isacharequal}\ {\isadigit{0}}\isanewline +\ \ \ \ \ \ \ \ then\ pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ k{\isacharparenright}\ x\isanewline +\ \ \ \ \ \ \ \ else\ {\isacharparenleft}pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ {\isacharparenleft}{\isacharminus}\ k{\isacharparenright}{\isacharparenright}\ x{\isacharparenright}{\isasymdiv}{\isacharparenright}{\isachardoublequoteclose}\isanewline +\isanewline +\ \ \ \ \isacommand{definition}\isamarkupfalse% +\isanewline +\ \ \ \ \ \ example\ {\isacharcolon}{\isacharcolon}\ int\isanewline +\ \ \ \ \ \ {\isachardoublequoteopen}example\ {\isacharequal}\ pow{\isacharunderscore}int\ {\isadigit{1}}{\isadigit{0}}\ {\isacharparenleft}{\isacharminus}{\isadigit{2}}{\isacharparenright}{\isachardoublequoteclose}% +\begin{isamarkuptext}% +\noindent Now we generate and compile code for SML:% +\end{isamarkuptext}% +\isamarkuptrue% +\ \ \ \ \isacommand{code{\isacharunderscore}gen}\isamarkupfalse% +\ example\ {\isacharparenleft}SML\ {\isacharminus}{\isacharparenright}% +\begin{isamarkuptext}% +\noindent The result is as expected:% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimML +\ \ \ \ % +\endisadelimML +% +\isatagML +\isacommand{ML}\isamarkupfalse% +\ {\isacharverbatimopen}\isanewline +\ \ \ \ \ \ if\ ROOT{\isachardot}Classes{\isachardot}example\ {\isacharequal}\ {\isachartilde}{\isadigit{2}}{\isadigit{0}}\ then\ {\isacharparenleft}{\isacharparenright}\ else\ error\ {\isachardoublequote}Wrong\ result{\isachardoublequote}\isanewline +\ \ \ \ {\isacharverbatimclose}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +\isanewline +% +\isadelimtheory +\isanewline +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 75c8a52f8447 -r 1df105407f87 doc-src/IsarAdvanced/Classes/Thy/document/isabelle.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarAdvanced/Classes/Thy/document/isabelle.sty Wed Oct 11 08:57:47 2006 +0200 @@ -0,0 +1,214 @@ +%% +%% +%% +%% macros for Isabelle generated LaTeX output +%% + +%%% Simple document preparation (based on theory token language and symbols) + +% isabelle environments + +\newcommand{\isabellecontext}{UNKNOWN} + +\newcommand{\isastyle}{\UNDEF} +\newcommand{\isastyleminor}{\UNDEF} +\newcommand{\isastylescript}{\UNDEF} +\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{\isascriptstyle${}\sb{#1}$}} +\newcommand{\isactrlsup}[1]{\emph{\isascriptstyle${}\sp{#1}$}} +\newcommand{\isactrlisub}[1]{\emph{\isascriptstyle${}\sb{#1}$}} +\newcommand{\isactrlisup}[1]{\emph{\isascriptstyle${}\sp{#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}} +\newcommand{\isactrlloc}[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}{} +\newcommand{\isadigit}[1]{#1} + +\newcommand{\isachardefaults}{% +\chardef\isacharbang=`\!% +\chardef\isachardoublequote=`\"% +\chardef\isachardoublequoteopen=`\"% +\chardef\isachardoublequoteclose=`\"% +\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=`\_% +\def\isacharunderscorekeyword{\_}% +\chardef\isacharbackquote=`\`% +\chardef\isacharbackquoteopen=`\`% +\chardef\isacharbackquoteclose=`\`% +\chardef\isacharbraceleft=`\{% +\chardef\isacharbar=`\|% +\chardef\isacharbraceright=`\}% +\chardef\isachartilde=`\~% +\def\isacharverbatimopen{\isacharbraceleft\isacharasterisk}% +\def\isacharverbatimclose{\isacharasterisk\isacharbraceright}% +} + + +% keyword and section markup + +\newcommand{\isakeyword}[1] +{\emph{\bf\def\isachardot{.}\def\isacharunderscore{\isacharunderscorekeyword}% +\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}{\par\isastyletext\begin{isapar}}{\end{isapar}} +\newenvironment{isamarkuptxt}{\par\isastyletxt\begin{isapar}}{\end{isapar}} +\newcommand{\isamarkupcmt}[1]{{\isastylecmt--- #1}} + + +% styles + +\def\isabellestyle#1{\csname isabellestyle#1\endcsname} + +\newcommand{\isabellestyledefault}{% +\renewcommand{\isastyle}{\small\tt\slshape}% +\renewcommand{\isastyleminor}{\small\tt\slshape}% +\renewcommand{\isastylescript}{\footnotesize\tt\slshape}% +\isachardefaults% +} +\isabellestyledefault + +\newcommand{\isabellestylett}{% +\renewcommand{\isastyle}{\small\tt}% +\renewcommand{\isastyleminor}{\small\tt}% +\renewcommand{\isastylescript}{\footnotesize\tt}% +\isachardefaults% +} + +\newcommand{\isabellestyleit}{% +\renewcommand{\isastyle}{\small\it}% +\renewcommand{\isastyleminor}{\it}% +\renewcommand{\isastylescript}{\footnotesize\it}% +\renewcommand{\isacharunderscorekeyword}{\mbox{-}}% +\renewcommand{\isacharbang}{\isamath{!}}% +\renewcommand{\isachardoublequote}{}% +\renewcommand{\isachardoublequoteopen}{}% +\renewcommand{\isachardoublequoteclose}{}% +\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}}}% +\renewcommand{\isacharbackquoteopen}{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\langle$}}}% +\renewcommand{\isacharbackquoteclose}{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\rangle$}}}% +\renewcommand{\isacharverbatimopen}{\isamath{\langle\!\langle}}% +\renewcommand{\isacharverbatimclose}{\isamath{\rangle\!\rangle}}% +} + +\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 75c8a52f8447 -r 1df105407f87 doc-src/IsarAdvanced/Classes/Thy/document/isabellesym.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarAdvanced/Classes/Thy/document/isabellesym.sty Wed Oct 11 08:57:47 2006 +0200 @@ -0,0 +1,360 @@ +%% +%% +%% +%% definitions of standard Isabelle symbols +%% + +\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{\isasymleadsto}{\isamath{\leadsto}} %requires amssymb +\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{\isasymColon}{\isamath{\mathrel{::}}} +\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{\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{\isasymnot}{\isamath{\neg}} +\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{\isasymsetminus}{\isamath{\setminus}} +\newcommand{\isasympropto}{\isamath{\propto}} +\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{\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}} +\newcommand{\isasymsome}{\isamath{\epsilon\,}} diff -r 75c8a52f8447 -r 1df105407f87 doc-src/IsarAdvanced/Classes/Thy/document/pdfsetup.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarAdvanced/Classes/Thy/document/pdfsetup.sty Wed Oct 11 08:57:47 2006 +0200 @@ -0,0 +1,25 @@ +%% +%% +%% +%% smart url or hyperref setup +%% + +\newif\ifpdfoutput +\ifx\pdfoutput\undefined +\else + \ifx\pdfoutput\relax + \else + \ifcase\pdfoutput + \else\pdfoutputtrue\fi + \fi +\fi + +\ifpdfoutput + \message{PDF output} + \usepackage{color}\definecolor{darkblue}{rgb}{0,0,0.5} + \usepackage[pdftex,a4paper,colorlinks=true,linkcolor=darkblue,citecolor=darkblue,filecolor=darkblue,pagecolor=darkblue,urlcolor=darkblue]{hyperref} + \IfFileExists{thumbpdf.sty}{\usepackage{thumbpdf}}{} +\else + \message{No PDF output} + \usepackage{url} +\fi diff -r 75c8a52f8447 -r 1df105407f87 doc-src/IsarAdvanced/Codegen/Thy/document/Codegen.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarAdvanced/Codegen/Thy/document/Codegen.tex Wed Oct 11 08:57:47 2006 +0200 @@ -0,0 +1,99 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Codegen}% +% +\isadelimtheory +\isanewline +\isanewline +\isanewline +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Codegen\isanewline +\isakeyword{imports}\ Main\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupchapter{Code generation from Isabelle theories% +} +\isamarkuptrue% +% +\isamarkupsection{Introduction% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\cite{isa-tutorial}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Basics% +} +\isamarkuptrue% +% +\isamarkupsubsection{Invoking the code generator% +} +\isamarkuptrue% +% +\isamarkupsubsection{Theorem selection% +} +\isamarkuptrue% +% +\isamarkupsubsection{Type classes% +} +\isamarkuptrue% +% +\isamarkupsubsection{Preprocessing% +} +\isamarkuptrue% +% +\isamarkupsubsection{Incremental code generation% +} +\isamarkuptrue% +% +\isamarkupsection{Customizing serialization% +} +\isamarkuptrue% +% +\isamarkupsection{Recipes and advanced topics% +} +\isamarkuptrue% +% +\isamarkupsection{ML interfaces% +} +\isamarkuptrue% +% +\isamarkupsubsection{codegen\_data.ML% +} +\isamarkuptrue% +% +\isamarkupsubsection{Implementing code generator applications% +} +\isamarkuptrue% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r 75c8a52f8447 -r 1df105407f87 doc-src/IsarAdvanced/Codegen/Thy/document/isabelle.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarAdvanced/Codegen/Thy/document/isabelle.sty Wed Oct 11 08:57:47 2006 +0200 @@ -0,0 +1,214 @@ +%% +%% +%% +%% macros for Isabelle generated LaTeX output +%% + +%%% Simple document preparation (based on theory token language and symbols) + +% isabelle environments + +\newcommand{\isabellecontext}{UNKNOWN} + +\newcommand{\isastyle}{\UNDEF} +\newcommand{\isastyleminor}{\UNDEF} +\newcommand{\isastylescript}{\UNDEF} +\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{\isascriptstyle${}\sb{#1}$}} +\newcommand{\isactrlsup}[1]{\emph{\isascriptstyle${}\sp{#1}$}} +\newcommand{\isactrlisub}[1]{\emph{\isascriptstyle${}\sb{#1}$}} +\newcommand{\isactrlisup}[1]{\emph{\isascriptstyle${}\sp{#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}} +\newcommand{\isactrlloc}[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}{} +\newcommand{\isadigit}[1]{#1} + +\newcommand{\isachardefaults}{% +\chardef\isacharbang=`\!% +\chardef\isachardoublequote=`\"% +\chardef\isachardoublequoteopen=`\"% +\chardef\isachardoublequoteclose=`\"% +\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=`\_% +\def\isacharunderscorekeyword{\_}% +\chardef\isacharbackquote=`\`% +\chardef\isacharbackquoteopen=`\`% +\chardef\isacharbackquoteclose=`\`% +\chardef\isacharbraceleft=`\{% +\chardef\isacharbar=`\|% +\chardef\isacharbraceright=`\}% +\chardef\isachartilde=`\~% +\def\isacharverbatimopen{\isacharbraceleft\isacharasterisk}% +\def\isacharverbatimclose{\isacharasterisk\isacharbraceright}% +} + + +% keyword and section markup + +\newcommand{\isakeyword}[1] +{\emph{\bf\def\isachardot{.}\def\isacharunderscore{\isacharunderscorekeyword}% +\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}{\par\isastyletext\begin{isapar}}{\end{isapar}} +\newenvironment{isamarkuptxt}{\par\isastyletxt\begin{isapar}}{\end{isapar}} +\newcommand{\isamarkupcmt}[1]{{\isastylecmt--- #1}} + + +% styles + +\def\isabellestyle#1{\csname isabellestyle#1\endcsname} + +\newcommand{\isabellestyledefault}{% +\renewcommand{\isastyle}{\small\tt\slshape}% +\renewcommand{\isastyleminor}{\small\tt\slshape}% +\renewcommand{\isastylescript}{\footnotesize\tt\slshape}% +\isachardefaults% +} +\isabellestyledefault + +\newcommand{\isabellestylett}{% +\renewcommand{\isastyle}{\small\tt}% +\renewcommand{\isastyleminor}{\small\tt}% +\renewcommand{\isastylescript}{\footnotesize\tt}% +\isachardefaults% +} + +\newcommand{\isabellestyleit}{% +\renewcommand{\isastyle}{\small\it}% +\renewcommand{\isastyleminor}{\it}% +\renewcommand{\isastylescript}{\footnotesize\it}% +\renewcommand{\isacharunderscorekeyword}{\mbox{-}}% +\renewcommand{\isacharbang}{\isamath{!}}% +\renewcommand{\isachardoublequote}{}% +\renewcommand{\isachardoublequoteopen}{}% +\renewcommand{\isachardoublequoteclose}{}% +\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}}}% +\renewcommand{\isacharbackquoteopen}{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\langle$}}}% +\renewcommand{\isacharbackquoteclose}{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\rangle$}}}% +\renewcommand{\isacharverbatimopen}{\isamath{\langle\!\langle}}% +\renewcommand{\isacharverbatimclose}{\isamath{\rangle\!\rangle}}% +} + +\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 75c8a52f8447 -r 1df105407f87 doc-src/IsarAdvanced/Codegen/Thy/document/isabellesym.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarAdvanced/Codegen/Thy/document/isabellesym.sty Wed Oct 11 08:57:47 2006 +0200 @@ -0,0 +1,360 @@ +%% +%% +%% +%% definitions of standard Isabelle symbols +%% + +\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{\isasymleadsto}{\isamath{\leadsto}} %requires amssymb +\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{\isasymColon}{\isamath{\mathrel{::}}} +\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{\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{\isasymnot}{\isamath{\neg}} +\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{\isasymsetminus}{\isamath{\setminus}} +\newcommand{\isasympropto}{\isamath{\propto}} +\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{\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}} +\newcommand{\isasymsome}{\isamath{\epsilon\,}} diff -r 75c8a52f8447 -r 1df105407f87 doc-src/IsarAdvanced/Codegen/Thy/document/pdfsetup.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarAdvanced/Codegen/Thy/document/pdfsetup.sty Wed Oct 11 08:57:47 2006 +0200 @@ -0,0 +1,25 @@ +%% +%% +%% +%% smart url or hyperref setup +%% + +\newif\ifpdfoutput +\ifx\pdfoutput\undefined +\else + \ifx\pdfoutput\relax + \else + \ifcase\pdfoutput + \else\pdfoutputtrue\fi + \fi +\fi + +\ifpdfoutput + \message{PDF output} + \usepackage{color}\definecolor{darkblue}{rgb}{0,0,0.5} + \usepackage[pdftex,a4paper,colorlinks=true,linkcolor=darkblue,citecolor=darkblue,filecolor=darkblue,pagecolor=darkblue,urlcolor=darkblue]{hyperref} + \IfFileExists{thumbpdf.sty}{\usepackage{thumbpdf}}{} +\else + \message{No PDF output} + \usepackage{url} +\fi