| author | wenzelm | 
| Fri, 18 Aug 2017 13:55:05 +0200 | |
| changeset 66452 | 450cefec7c11 | 
| parent 64525 | 9c3da2276e19 | 
| child 67139 | 8fe0aba577af | 
| permissions | -rw-r--r-- | 
| 7709 | 1  | 
%%  | 
2  | 
%% macros for Isabelle generated LaTeX output  | 
|
3  | 
%%  | 
|
4  | 
||
| 9913 | 5  | 
%%% Simple document preparation (based on theory token language and symbols)  | 
| 7709 | 6  | 
|
| 8714 | 7  | 
% isabelle environments  | 
| 7788 | 8  | 
|
| 9913 | 9  | 
\newcommand{\isabellecontext}{UNKNOWN}
 | 
| 
58870
 
e2c0d8ef29cb
more flexibile \setisabellecontext, independently of header;
 
wenzelm 
parents: 
58868 
diff
changeset
 | 
10  | 
\newcommand{\setisabellecontext}[1]{\def\isabellecontext{#1}}
 | 
| 9913 | 11  | 
|
| 
18860
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
12  | 
\newcommand{\isastyle}{\UNDEF}
 | 
| 
58716
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
13  | 
\newcommand{\isastylett}{\UNDEF}
 | 
| 
18860
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
14  | 
\newcommand{\isastyleminor}{\UNDEF}
 | 
| 
58716
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
15  | 
\newcommand{\isastyleminortt}{\UNDEF}
 | 
| 
18860
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
16  | 
\newcommand{\isastylescript}{\UNDEF}
 | 
| 9657 | 17  | 
\newcommand{\isastyletext}{\normalsize\rm}
 | 
18  | 
\newcommand{\isastyletxt}{\rm}
 | 
|
19  | 
\newcommand{\isastylecmt}{\rm}
 | 
|
| 8713 | 20  | 
|
| 43327 | 21  | 
\newcommand{\isaspacing}{%
 | 
22  | 
\sfcode 42 1000 % .  | 
|
23  | 
\sfcode 63 1000 % ?  | 
|
24  | 
\sfcode 33 1000 % !  | 
|
25  | 
\sfcode 58 1000 % :  | 
|
26  | 
\sfcode 59 1000 % ;  | 
|
27  | 
\sfcode 44 1000 % ,  | 
|
28  | 
}  | 
|
| 
43325
 
4384f4ae0574
some attempts at robust \isaspacing so that \isa{...} can be used in section headings etc. (need to avoid `\? for some reason);
 
wenzelm 
parents: 
43321 
diff
changeset
 | 
29  | 
|
| 10220 | 30  | 
%symbol markup -- \emph achieves decent spacing via italic corrections  | 
31  | 
\newcommand{\isamath}[1]{\emph{$#1$}}
 | 
|
32  | 
\newcommand{\isatext}[1]{\emph{#1}}
 | 
|
| 
43325
 
4384f4ae0574
some attempts at robust \isaspacing so that \isa{...} can be used in section headings etc. (need to avoid `\? for some reason);
 
wenzelm 
parents: 
43321 
diff
changeset
 | 
33  | 
\DeclareRobustCommand{\isascriptstyle}{\def\isamath##1{##1}\def\isatext##1{\mbox{\isaspacing\isastylescript##1}}}
 | 
| 
17180
 
5fefe658a6f8
recover original definitions of \isactrlsub etc.;
 
wenzelm 
parents: 
17173 
diff
changeset
 | 
34  | 
\newcommand{\isactrlsub}[1]{\emph{\isascriptstyle${}\sb{#1}$}}
 | 
| 
 
5fefe658a6f8
recover original definitions of \isactrlsub etc.;
 
wenzelm 
parents: 
17173 
diff
changeset
 | 
35  | 
\newcommand{\isactrlsup}[1]{\emph{\isascriptstyle${}\sp{#1}$}}
 | 
| 
43325
 
4384f4ae0574
some attempts at robust \isaspacing so that \isa{...} can be used in section headings etc. (need to avoid `\? for some reason);
 
wenzelm 
parents: 
43321 
diff
changeset
 | 
36  | 
\DeclareRobustCommand{\isactrlbsub}{\emph\bgroup\math{}\sb\bgroup\mbox\bgroup\isaspacing\isastylescript}
 | 
| 25081 | 37  | 
\DeclareRobustCommand{\isactrlesub}{\egroup\egroup\endmath\egroup}
 | 
| 
43325
 
4384f4ae0574
some attempts at robust \isaspacing so that \isa{...} can be used in section headings etc. (need to avoid `\? for some reason);
 
wenzelm 
parents: 
43321 
diff
changeset
 | 
38  | 
\DeclareRobustCommand{\isactrlbsup}{\emph\bgroup\math{}\sp\bgroup\mbox\bgroup\isaspacing\isastylescript}
 | 
| 25081 | 39  | 
\DeclareRobustCommand{\isactrlesup}{\egroup\egroup\endmath\egroup}
 | 
| 11573 | 40  | 
\newcommand{\isactrlbold}[1]{{\bfseries\upshape\boldmath#1}}
 | 
| 27337 | 41  | 
|
| 
61473
 
34d1913f0b20
clarified control antiquotations: decode control symbol to get name;
 
wenzelm 
parents: 
61463 
diff
changeset
 | 
42  | 
\newcommand{\isaantiqcontrol}[1]{\isatt{{\char`\\}{\char`\<}{\char`\^}#1{\char`\>}}}
 | 
| 22647 | 43  | 
\newenvironment{isaantiq}{{\isacharat\isacharbraceleft}}{{\isacharbraceright}}
 | 
| 
14332
 
fd3535af90ab
spanning super and sub scripts \<^bsub> .. \<^esub> and \<^bsup> .. \<^esup>
 
kleing 
parents: 
14234 
diff
changeset
 | 
44  | 
|
| 8713 | 45  | 
\newdimen\isa@parindent\newdimen\isa@parskip  | 
| 9657 | 46  | 
|
| 
9702
 
f23bee3c0682
\newenvironment{isabellebody}: version without trivlist;
 
wenzelm 
parents: 
9688 
diff
changeset
 | 
47  | 
\newenvironment{isabellebody}{%
 | 
| 
11863
 
87643169ae7d
\newif\ifisamarkup controls spacing of isabeginpar;
 
wenzelm 
parents: 
11573 
diff
changeset
 | 
48  | 
\isamarkuptrue\par%  | 
| 
9702
 
f23bee3c0682
\newenvironment{isabellebody}: version without trivlist;
 
wenzelm 
parents: 
9688 
diff
changeset
 | 
49  | 
\isa@parindent\parindent\parindent0pt%  | 
| 8713 | 50  | 
\isa@parskip\parskip\parskip0pt%  | 
| 
43325
 
4384f4ae0574
some attempts at robust \isaspacing so that \isa{...} can be used in section headings etc. (need to avoid `\? for some reason);
 
wenzelm 
parents: 
43321 
diff
changeset
 | 
51  | 
\isaspacing\isastyle}{\par}
 | 
| 
9702
 
f23bee3c0682
\newenvironment{isabellebody}: version without trivlist;
 
wenzelm 
parents: 
9688 
diff
changeset
 | 
52  | 
|
| 
58716
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
53  | 
\newenvironment{isabellebodytt}{%
 | 
| 
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
54  | 
\isamarkuptrue\par%  | 
| 
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
55  | 
\isa@parindent\parindent\parindent0pt%  | 
| 
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
56  | 
\isa@parskip\parskip\parskip0pt%  | 
| 
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
57  | 
\isaspacing\isastylett}{\par}
 | 
| 
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
58  | 
|
| 
9702
 
f23bee3c0682
\newenvironment{isabellebody}: version without trivlist;
 
wenzelm 
parents: 
9688 
diff
changeset
 | 
59  | 
\newenvironment{isabelle}
 | 
| 10422 | 60  | 
{\begin{trivlist}\begin{isabellebody}\item\relax}
 | 
61  | 
{\end{isabellebody}\end{trivlist}}
 | 
|
| 8713 | 62  | 
|
| 
58716
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
63  | 
\newenvironment{isabellett}
 | 
| 
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
64  | 
{\begin{trivlist}\begin{isabellebodytt}\item\relax}
 | 
| 
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
65  | 
{\end{isabellebodytt}\end{trivlist}}
 | 
| 
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
66  | 
|
| 
43325
 
4384f4ae0574
some attempts at robust \isaspacing so that \isa{...} can be used in section headings etc. (need to avoid `\? for some reason);
 
wenzelm 
parents: 
43321 
diff
changeset
 | 
67  | 
\newcommand{\isa}[1]{\emph{\isaspacing\isastyleminor #1}}
 | 
| 
58716
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
68  | 
\newcommand{\isatt}[1]{\emph{\isaspacing\isastyleminortt #1}}
 | 
| 8714 | 69  | 
|
| 10949 | 70  | 
\newcommand{\isaindent}[1]{\hphantom{#1}}
 | 
| 13933 | 71  | 
\newcommand{\isanewline}{\mbox{}\par\mbox{}}
 | 
| 
17158
 
d68bf267cbba
added \isachardoublequoteopen/close, \isacharbackquoteopen/close;
 
wenzelm 
parents: 
17052 
diff
changeset
 | 
72  | 
\newcommand{\isasep}{}
 | 
| 9669 | 73  | 
\newcommand{\isadigit}[1]{#1}
 | 
| 7752 | 74  | 
|
| 
18860
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
75  | 
\newcommand{\isachardefaults}{%
 | 
| 
64525
 
9c3da2276e19
avoid extra space intruding rail diagrams (amending 4854f7ee0987);
 
wenzelm 
parents: 
63590 
diff
changeset
 | 
76  | 
\def\isacharbell{\isamath{\bigbox}}%requires stmaryrd
 | 
| 
18860
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
77  | 
\chardef\isacharbang=`\!%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
78  | 
\chardef\isachardoublequote=`\"%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
79  | 
\chardef\isachardoublequoteopen=`\"%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
80  | 
\chardef\isachardoublequoteclose=`\"%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
81  | 
\chardef\isacharhash=`\#%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
82  | 
\chardef\isachardollar=`\$%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
83  | 
\chardef\isacharpercent=`\%%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
84  | 
\chardef\isacharampersand=`\&%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
85  | 
\chardef\isacharprime=`\'%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
86  | 
\chardef\isacharparenleft=`\(%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
87  | 
\chardef\isacharparenright=`\)%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
88  | 
\chardef\isacharasterisk=`\*%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
89  | 
\chardef\isacharplus=`\+%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
90  | 
\chardef\isacharcomma=`\,%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
91  | 
\chardef\isacharminus=`\-%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
92  | 
\chardef\isachardot=`\.%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
93  | 
\chardef\isacharslash=`\/%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
94  | 
\chardef\isacharcolon=`\:%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
95  | 
\chardef\isacharsemicolon=`\;%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
96  | 
\chardef\isacharless=`\<%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
97  | 
\chardef\isacharequal=`\=%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
98  | 
\chardef\isachargreater=`\>%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
99  | 
\chardef\isacharquery=`\?%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
100  | 
\chardef\isacharat=`\@%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
101  | 
\chardef\isacharbrackleft=`\[%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
102  | 
\chardef\isacharbackslash=`\\%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
103  | 
\chardef\isacharbrackright=`\]%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
104  | 
\chardef\isacharcircum=`\^%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
105  | 
\chardef\isacharunderscore=`\_%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
106  | 
\def\isacharunderscorekeyword{\_}%
 | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
107  | 
\chardef\isacharbackquote=`\`%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
108  | 
\chardef\isacharbackquoteopen=`\`%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
109  | 
\chardef\isacharbackquoteclose=`\`%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
110  | 
\chardef\isacharbraceleft=`\{%
 | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
111  | 
\chardef\isacharbar=`\|%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
112  | 
\chardef\isacharbraceright=`\}%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
113  | 
\chardef\isachartilde=`\~%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
114  | 
\def\isacharverbatimopen{\isacharbraceleft\isacharasterisk}%
 | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
115  | 
\def\isacharverbatimclose{\isacharasterisk\isacharbraceright}%
 | 
| 55033 | 116  | 
\def\isacartoucheopen{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\langle$}}}%
 | 
117  | 
\def\isacartoucheclose{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\rangle$}}}%
 | 
|
| 
18860
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
118  | 
}  | 
| 7788 | 119  | 
|
| 7797 | 120  | 
|
121  | 
% keyword and section markup  | 
|
| 7788 | 122  | 
|
| 9669 | 123  | 
\newcommand{\isakeyword}[1]
 | 
| 
18860
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
124  | 
{\emph{\bf\def\isachardot{.}\def\isacharunderscore{\isacharunderscorekeyword}%
 | 
| 9669 | 125  | 
\def\isacharbraceleft{\{}\def\isacharbraceright{\}}#1}}
 | 
126  | 
\newcommand{\isacommand}[1]{\isakeyword{#1}}
 | 
|
| 7758 | 127  | 
|
| 8501 | 128  | 
\newcommand{\isamarkupheader}[1]{\section{#1}}
 | 
| 7752 | 129  | 
\newcommand{\isamarkupchapter}[1]{\chapter{#1}}
 | 
130  | 
\newcommand{\isamarkupsection}[1]{\section{#1}}
 | 
|
131  | 
\newcommand{\isamarkupsubsection}[1]{\subsection{#1}}
 | 
|
132  | 
\newcommand{\isamarkupsubsubsection}[1]{\subsubsection{#1}}
 | 
|
| 61463 | 133  | 
\newcommand{\isamarkupparagraph}[1]{\paragraph{#1}}
 | 
134  | 
\newcommand{\isamarkupsubparagraph}[1]{\subparagraph{#1}}
 | 
|
| 
8662
 
f9679ddbc492
isapar, isamarkuptext, isamarkuptxt turned into environments;
 
wenzelm 
parents: 
8512 
diff
changeset
 | 
135  | 
|
| 
11863
 
87643169ae7d
\newif\ifisamarkup controls spacing of isabeginpar;
 
wenzelm 
parents: 
11573 
diff
changeset
 | 
136  | 
\newif\ifisamarkup  | 
| 
 
87643169ae7d
\newif\ifisamarkup controls spacing of isabeginpar;
 
wenzelm 
parents: 
11573 
diff
changeset
 | 
137  | 
\newcommand{\isabeginpar}{\par\ifisamarkup\relax\else\medskip\fi}
 | 
| 10586 | 138  | 
\newcommand{\isaendpar}{\par\medskip}
 | 
139  | 
\newenvironment{isapar}{\parindent\isa@parindent\parskip\isa@parskip\isabeginpar}{\isaendpar}
 | 
|
| 
17215
 
8b969275a5d2
isamarkuptext/txt: \par before changing sizes prevents spacing anomaly;
 
wenzelm 
parents: 
17180 
diff
changeset
 | 
140  | 
\newenvironment{isamarkuptext}{\par\isastyletext\begin{isapar}}{\end{isapar}}
 | 
| 
 
8b969275a5d2
isamarkuptext/txt: \par before changing sizes prevents spacing anomaly;
 
wenzelm 
parents: 
17180 
diff
changeset
 | 
141  | 
\newenvironment{isamarkuptxt}{\par\isastyletxt\begin{isapar}}{\end{isapar}}
 | 
| 9657 | 142  | 
\newcommand{\isamarkupcmt}[1]{{\isastylecmt--- #1}}
 | 
143  | 
||
144  | 
||
| 
18860
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
145  | 
% styles  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
146  | 
|
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
147  | 
\def\isabellestyle#1{\csname isabellestyle#1\endcsname}
 | 
| 9657 | 148  | 
|
| 
18860
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
149  | 
\newcommand{\isabellestyledefault}{%
 | 
| 42514 | 150  | 
\def\isastyle{\small\tt\slshape}%
 | 
| 
58716
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
151  | 
\def\isastylett{\small\tt}%
 | 
| 42514 | 152  | 
\def\isastyleminor{\small\tt\slshape}%
 | 
| 
58716
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
153  | 
\def\isastyleminortt{\small\tt}%
 | 
| 42514 | 154  | 
\def\isastylescript{\footnotesize\tt\slshape}%
 | 
| 
18860
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
155  | 
\isachardefaults%  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
156  | 
}  | 
| 
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
157  | 
\isabellestyledefault  | 
| 9657 | 158  | 
|
| 10024 | 159  | 
\newcommand{\isabellestylett}{%
 | 
| 42514 | 160  | 
\def\isastyle{\small\tt}%
 | 
| 
58716
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
161  | 
\def\isastylett{\small\tt}%
 | 
| 42514 | 162  | 
\def\isastyleminor{\small\tt}%
 | 
| 
58716
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
163  | 
\def\isastyleminortt{\small\tt}%
 | 
| 42514 | 164  | 
\def\isastylescript{\footnotesize\tt}%
 | 
| 
18860
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
165  | 
\isachardefaults%  | 
| 10024 | 166  | 
}  | 
| 
18860
 
9089cdb4c5fd
all styles now reset to defaults first, i.e. the document may switch styles back and forth;
 
wenzelm 
parents: 
17531 
diff
changeset
 | 
167  | 
|
| 9669 | 168  | 
\newcommand{\isabellestyleit}{%
 | 
| 42514 | 169  | 
\def\isastyle{\small\it}%
 | 
| 
58716
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
170  | 
\def\isastylett{\small\tt}%
 | 
| 42514 | 171  | 
\def\isastyleminor{\it}%
 | 
| 
58716
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
172  | 
\def\isastyleminortt{\tt}%
 | 
| 42514 | 173  | 
\def\isastylescript{\footnotesize\it}%
 | 
174  | 
\isachardefaults%  | 
|
175  | 
\def\isacharunderscorekeyword{\mbox{-}}%
 | 
|
176  | 
\def\isacharbang{\isamath{!}}%
 | 
|
| 
49320
 
94bd2fb83d11
discontinued experiment with literal replacement text in PDF (cf. b646316f8b3c, 2ff10e613689);
 
wenzelm 
parents: 
45646 
diff
changeset
 | 
177  | 
\def\isachardoublequote{}%
 | 
| 
 
94bd2fb83d11
discontinued experiment with literal replacement text in PDF (cf. b646316f8b3c, 2ff10e613689);
 
wenzelm 
parents: 
45646 
diff
changeset
 | 
178  | 
\def\isachardoublequoteopen{}%
 | 
| 
 
94bd2fb83d11
discontinued experiment with literal replacement text in PDF (cf. b646316f8b3c, 2ff10e613689);
 
wenzelm 
parents: 
45646 
diff
changeset
 | 
179  | 
\def\isachardoublequoteclose{}%
 | 
| 42514 | 180  | 
\def\isacharhash{\isamath{\#}}%
 | 
181  | 
\def\isachardollar{\isamath{\$}}%
 | 
|
182  | 
\def\isacharpercent{\isamath{\%}}%
 | 
|
183  | 
\def\isacharampersand{\isamath{\&}}%
 | 
|
184  | 
\def\isacharprime{\isamath{\mskip2mu{'}\mskip-2mu}}%
 | 
|
185  | 
\def\isacharparenleft{\isamath{(}}%
 | 
|
186  | 
\def\isacharparenright{\isamath{)}}%
 | 
|
187  | 
\def\isacharasterisk{\isamath{*}}%
 | 
|
188  | 
\def\isacharplus{\isamath{+}}%
 | 
|
189  | 
\def\isacharcomma{\isamath{\mathord,}}%
 | 
|
190  | 
\def\isacharminus{\isamath{-}}%
 | 
|
191  | 
\def\isachardot{\isamath{\mathord.}}%
 | 
|
192  | 
\def\isacharslash{\isamath{/}}%
 | 
|
193  | 
\def\isacharcolon{\isamath{\mathord:}}%
 | 
|
194  | 
\def\isacharsemicolon{\isamath{\mathord;}}%
 | 
|
195  | 
\def\isacharless{\isamath{<}}%
 | 
|
196  | 
\def\isacharequal{\isamath{=}}%
 | 
|
197  | 
\def\isachargreater{\isamath{>}}%
 | 
|
198  | 
\def\isacharat{\isamath{@}}%
 | 
|
199  | 
\def\isacharbrackleft{\isamath{[}}%
 | 
|
200  | 
\def\isacharbackslash{\isamath{\backslash}}%
 | 
|
201  | 
\def\isacharbrackright{\isamath{]}}%
 | 
|
202  | 
\def\isacharunderscore{\mbox{-}}%
 | 
|
203  | 
\def\isacharbraceleft{\isamath{\{}}%
 | 
|
204  | 
\def\isacharbar{\isamath{\mid}}%
 | 
|
205  | 
\def\isacharbraceright{\isamath{\}}}%
 | 
|
206  | 
\def\isachartilde{\isamath{{}\sp{\sim}}}%
 | 
|
207  | 
\def\isacharbackquoteopen{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\langle$}}}%
 | 
|
208  | 
\def\isacharbackquoteclose{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\rangle$}}}%
 | 
|
209  | 
\def\isacharverbatimopen{\isamath{\langle\!\langle}}%
 | 
|
210  | 
\def\isacharverbatimclose{\isamath{\rangle\!\rangle}}%
 | 
|
| 9657 | 211  | 
}  | 
212  | 
||
| 
45646
 
02afa20cf397
refined "literal" document style, with some correspondence to actual text source;
 
wenzelm 
parents: 
43458 
diff
changeset
 | 
213  | 
\newcommand{\isabellestyleliteral}{%
 | 
| 
42663
 
c31df4184ead
provide \isabellestyle{itunderscore} (requires underscore.sty);
 
wenzelm 
parents: 
42514 
diff
changeset
 | 
214  | 
\isabellestyleit%  | 
| 
 
c31df4184ead
provide \isabellestyle{itunderscore} (requires underscore.sty);
 
wenzelm 
parents: 
42514 
diff
changeset
 | 
215  | 
\def\isacharunderscore{\_}%
 | 
| 
 
c31df4184ead
provide \isabellestyle{itunderscore} (requires underscore.sty);
 
wenzelm 
parents: 
42514 
diff
changeset
 | 
216  | 
\def\isacharunderscorekeyword{\_}%
 | 
| 
45646
 
02afa20cf397
refined "literal" document style, with some correspondence to actual text source;
 
wenzelm 
parents: 
43458 
diff
changeset
 | 
217  | 
\chardef\isacharbackquoteopen=`\`%  | 
| 
 
02afa20cf397
refined "literal" document style, with some correspondence to actual text source;
 
wenzelm 
parents: 
43458 
diff
changeset
 | 
218  | 
\chardef\isacharbackquoteclose=`\`%  | 
| 
42663
 
c31df4184ead
provide \isabellestyle{itunderscore} (requires underscore.sty);
 
wenzelm 
parents: 
42514 
diff
changeset
 | 
219  | 
}  | 
| 
 
c31df4184ead
provide \isabellestyle{itunderscore} (requires underscore.sty);
 
wenzelm 
parents: 
42514 
diff
changeset
 | 
220  | 
|
| 55365 | 221  | 
\newcommand{\isabellestyleliteralunderscore}{%
 | 
222  | 
\isabellestyleliteral%  | 
|
223  | 
\def\isacharunderscore{\textunderscore}%
 | 
|
224  | 
\def\isacharunderscorekeyword{\textunderscore}%
 | 
|
225  | 
}  | 
|
226  | 
||
| 9975 | 227  | 
\newcommand{\isabellestylesl}{%
 | 
228  | 
\isabellestyleit%  | 
|
| 42514 | 229  | 
\def\isastyle{\small\sl}%
 | 
| 
58716
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
230  | 
\def\isastylett{\small\tt}%
 | 
| 42514 | 231  | 
\def\isastyleminor{\sl}%
 | 
| 
58716
 
23a380cc45f4
official support for "tt" style variants, avoid fragile \verb in LaTeX;
 
wenzelm 
parents: 
55365 
diff
changeset
 | 
232  | 
\def\isastyleminortt{\tt}%
 | 
| 42514 | 233  | 
\def\isastylescript{\footnotesize\sl}%
 | 
| 9975 | 234  | 
}  | 
| 17052 | 235  | 
|
236  | 
||
237  | 
% tagged regions  | 
|
238  | 
||
239  | 
%plain TeX version of comment package -- much faster!  | 
|
240  | 
\let\isafmtname\fmtname\def\fmtname{plain}
 | 
|
241  | 
\usepackage{comment}
 | 
|
242  | 
\let\fmtname\isafmtname  | 
|
243  | 
||
244  | 
\newcommand{\isafold}[1]{\emph{$\langle\mathord{\mathit{#1}}\rangle$}}
 | 
|
245  | 
||
246  | 
\newcommand{\isakeeptag}[1]%
 | 
|
247  | 
{\includecomment{isadelim#1}\includecomment{isatag#1}\csarg\def{isafold#1}{}}
 | 
|
248  | 
\newcommand{\isadroptag}[1]%
 | 
|
249  | 
{\excludecomment{isadelim#1}\excludecomment{isatag#1}\csarg\def{isafold#1}{}}
 | 
|
250  | 
\newcommand{\isafoldtag}[1]%
 | 
|
251  | 
{\includecomment{isadelim#1}\excludecomment{isatag#1}\csarg\def{isafold#1}{\isafold{#1}}}
 | 
|
252  | 
||
253  | 
\isakeeptag{theory}
 | 
|
254  | 
\isakeeptag{proof}
 | 
|
255  | 
\isakeeptag{ML}
 | 
|
256  | 
\isakeeptag{visible}
 | 
|
257  | 
\isadroptag{invisible}
 | 
|
258  | 
||
259  | 
\IfFileExists{isabelletags.sty}{\usepackage{isabelletags}}{}
 |