# HG changeset patch # User wenzelm # Date 1124441405 -7200 # Node ID e6a82d1a1829370308a313e8ed804645899913e2 # Parent 19ea4a0f4ec91eb4065a36863e4d0f98363b25ef updated; diff -r 19ea4a0f4ec9 -r e6a82d1a1829 doc-src/IsarOverview/Isar/document/Induction.tex --- a/doc-src/IsarOverview/Isar/document/Induction.tex Fri Aug 19 09:40:44 2005 +0200 +++ b/doc-src/IsarOverview/Isar/document/Induction.tex Fri Aug 19 10:50:05 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Induction}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsection{Case distinction and induction \label{sec:Induct}% } @@ -23,8 +36,14 @@ We have already met the \isa{cases} method for performing binary case splits. Here is another example:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isasymnot}\ A\ {\isasymor}\ A{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\ cases\isanewline \ \ \isamarkupfalse% @@ -38,7 +57,14 @@ \isacommand{thus}\ {\isacharquery}thesis\ \isamarkupfalse% \isacommand{{\isachardot}{\isachardot}}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent The two cases must come in this order because \isa{cases} merely abbreviates \isa{{\isacharparenleft}rule\ case{\isacharunderscore}split{\isacharunderscore}thm{\isacharparenright}} where @@ -52,8 +78,14 @@ However, if \isa{A} is large, we do not want to repeat it. This can be avoided by the following idiom% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isasymnot}\ A\ {\isasymor}\ A{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\ {\isacharparenleft}cases\ {\isachardoublequote}A{\isachardoublequote}{\isacharparenright}\isanewline \ \ \isamarkupfalse% @@ -67,7 +99,14 @@ \isacommand{thus}\ {\isacharquery}thesis\ \isamarkupfalse% \isacommand{{\isachardot}{\isachardot}}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent which is like the previous proof but instantiates @@ -80,9 +119,14 @@ where \isa{tl} is the tail of a list, and \isa{length} returns a natural number (remember: $0-1=0$):% \end{isamarkuptext}% -\isamarkuptrue% \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}length{\isacharparenleft}tl\ xs{\isacharparenright}\ {\isacharequal}\ length\ xs\ {\isacharminus}\ {\isadigit{1}}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\ {\isacharparenleft}cases\ xs{\isacharparenright}\isanewline \ \ \isamarkupfalse% @@ -96,7 +140,14 @@ \isacommand{thus}\ {\isacharquery}thesis\ \isamarkupfalse% \isacommand{by}\ simp\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Here `\isakeyword{case}~\isa{Nil}' abbreviates @@ -126,10 +177,23 @@ \begin{isamarkuptext}% We start with an inductive proof where both cases are proved automatically:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isadigit{2}}\ {\isacharasterisk}\ {\isacharparenleft}{\isasymSum}i{\isacharcolon}{\isacharcolon}nat{\isasymle}n{\isachardot}\ i{\isacharparenright}\ {\isacharequal}\ n{\isacharasterisk}{\isacharparenleft}n{\isacharplus}{\isadigit{1}}{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}induct\ n{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}\isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}induct\ n{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent The constraint \isa{{\isacharcolon}{\isacharcolon}nat} is needed because all of @@ -139,8 +203,14 @@ proof, we can use pattern matching to avoid having to repeat the goal statement:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isadigit{2}}\ {\isacharasterisk}\ {\isacharparenleft}{\isasymSum}i{\isacharcolon}{\isacharcolon}nat{\isasymle}n{\isachardot}\ i{\isacharparenright}\ {\isacharequal}\ n{\isacharasterisk}{\isacharparenleft}n{\isacharplus}{\isadigit{1}}{\isacharparenright}{\isachardoublequote}\ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequote}{\isacharquery}P\ n{\isachardoublequote}{\isacharparenright}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\ {\isacharparenleft}induct\ n{\isacharparenright}\isanewline \ \ \isamarkupfalse% @@ -155,15 +225,28 @@ \isacommand{thus}\ {\isachardoublequote}{\isacharquery}P{\isacharparenleft}Suc\ n{\isacharparenright}{\isachardoublequote}\ \isamarkupfalse% \isacommand{by}\ simp\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent We could refine this further to show more of the equational proof. Instead we explore the same avenue as for case distinctions: introducing context via the \isakeyword{case} command:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isadigit{2}}\ {\isacharasterisk}\ {\isacharparenleft}{\isasymSum}i{\isacharcolon}{\isacharcolon}nat\ {\isasymle}\ n{\isachardot}\ i{\isacharparenright}\ {\isacharequal}\ n{\isacharasterisk}{\isacharparenleft}n{\isacharplus}{\isadigit{1}}{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\ {\isacharparenleft}induct\ n{\isacharparenright}\isanewline \ \ \isamarkupfalse% @@ -177,7 +260,14 @@ \isacommand{thus}\ {\isacharquery}case\ \isamarkupfalse% \isacommand{by}\ simp\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent The implicitly defined \isa{{\isacharquery}case} refers to the @@ -189,8 +279,14 @@ (in contrast to the previous proof). The solution is the one outlined for \isa{Cons} above: replace \isa{Suc} by \isa{{\isacharparenleft}Suc\ i{\isacharparenright}}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ \isakeyword{fixes}\ n{\isacharcolon}{\isacharcolon}nat\ \isakeyword{shows}\ {\isachardoublequote}n\ {\isacharless}\ n{\isacharasterisk}n\ {\isacharplus}\ {\isadigit{1}}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\ {\isacharparenleft}induct\ n{\isacharparenright}\isanewline \ \ \isamarkupfalse% @@ -204,7 +300,14 @@ \isacommand{thus}\ {\isachardoublequote}Suc\ i\ {\isacharless}\ Suc\ i\ {\isacharasterisk}\ Suc\ i\ {\isacharplus}\ {\isadigit{1}}{\isachardoublequote}\ \isamarkupfalse% \isacommand{by}\ simp\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Of course we could again have written @@ -232,9 +335,15 @@ As an example we will now prove complete induction via structural induction.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ \isakeyword{assumes}\ A{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isasymAnd}n{\isachardot}\ {\isacharparenleft}{\isasymAnd}m{\isachardot}\ m\ {\isacharless}\ n\ {\isasymLongrightarrow}\ P\ m{\isacharparenright}\ {\isasymLongrightarrow}\ P\ n{\isacharparenright}{\isachardoublequote}\isanewline \ \ \isakeyword{shows}\ {\isachardoublequote}P{\isacharparenleft}n{\isacharcolon}{\isacharcolon}nat{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\ {\isacharparenleft}rule\ A{\isacharparenright}\isanewline \ \ \isamarkupfalse% @@ -285,7 +394,14 @@ \ \ \isamarkupfalse% \isacommand{qed}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Given the explanations above and the comments in the @@ -315,13 +431,13 @@ for details. As an example we define our own version of the reflexive transitive closure of a relation --- HOL provides a predefined one as well.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ rtc\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}set\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}set{\isachardoublequote}\ \ \ {\isacharparenleft}{\isachardoublequote}{\isacharunderscore}{\isacharasterisk}{\isachardoublequote}\ {\isacharbrackleft}{\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isadigit{0}}{\isacharbrackright}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{inductive}\ {\isachardoublequote}r{\isacharasterisk}{\isachardoublequote}\isanewline \isakeyword{intros}\isanewline refl{\isacharcolon}\ \ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}x{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isanewline -step{\isacharcolon}\ \ {\isachardoublequote}{\isasymlbrakk}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharsemicolon}\ {\isacharparenleft}y{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isamarkupfalse% +step{\isacharcolon}\ \ {\isachardoublequote}{\isasymlbrakk}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharsemicolon}\ {\isacharparenleft}y{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -329,8 +445,14 @@ relations (with concrete syntax \isa{r{\isacharasterisk}} instead of \isa{rtc\ r}), then the defining clauses are given. We will now prove that \isa{r{\isacharasterisk}} is indeed transitive:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ \isakeyword{assumes}\ A{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\ \isakeyword{shows}\ {\isachardoublequote}{\isacharparenleft}y{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{using}\ A\isanewline \isamarkupfalse% @@ -346,7 +468,14 @@ \isacommand{thus}\ {\isacharquery}case\ \isamarkupfalse% \isacommand{by}{\isacharparenleft}blast\ intro{\isacharcolon}\ rtc{\isachardot}step{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Rule induction is triggered by a fact $(x_1,\dots,x_n) @@ -357,9 +486,15 @@ However, this proof is rather terse. Here is a more readable version:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ \isakeyword{assumes}\ A{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\ \isakeyword{and}\ B{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}y{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isanewline \ \ \isakeyword{shows}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\ {\isacharminus}\isanewline \ \ \isamarkupfalse% @@ -391,7 +526,14 @@ \ \ \isamarkupfalse% \isacommand{qed}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent We start the proof with \isakeyword{from}~\isa{A\ B}. Only \isa{A} is ``consumed'' by the induction step. @@ -425,7 +567,7 @@ The example is an unusual definition of rotation:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ rot\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{recdef}\ rot\ {\isachardoublequote}measure\ length{\isachardoublequote}\ \ % @@ -434,7 +576,7 @@ \isanewline {\isachardoublequote}rot\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}\isanewline {\isachardoublequote}rot\ {\isacharbrackleft}x{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}x{\isacharbrackright}{\isachardoublequote}\isanewline -{\isachardoublequote}rot\ {\isacharparenleft}x{\isacharhash}y{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ y\ {\isacharhash}\ rot{\isacharparenleft}x{\isacharhash}zs{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +{\isachardoublequote}rot\ {\isacharparenleft}x{\isacharhash}y{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ y\ {\isacharhash}\ rot{\isacharparenleft}x{\isacharhash}zs{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent This yields, among other things, the induction rule @@ -446,8 +588,14 @@ only with datatypes and inductively defined sets, but not with recursive functions.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymLongrightarrow}\ rot\ xs\ {\isacharequal}\ tl\ xs\ {\isacharat}\ {\isacharbrackleft}hd\ xs{\isacharbrackright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\ {\isacharparenleft}induct\ xs\ rule{\isacharcolon}\ rot{\isachardot}induct{\isacharparenright}\isanewline \ \ \isamarkupfalse% @@ -480,7 +628,14 @@ \isacommand{show}\ {\isacharquery}case\ \isamarkupfalse% \isacommand{{\isachardot}}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -496,11 +651,34 @@ The proof is so simple that it can be condensed to% \end{isamarkuptext}% -\isamarkuptrue% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{by}\ {\isacharparenleft}induct\ xs\ rule{\isacharcolon}\ rot{\isachardot}induct{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}\isanewline -\isamarkupfalse% -\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r 19ea4a0f4ec9 -r e6a82d1a1829 doc-src/IsarOverview/Isar/document/Logic.tex --- a/doc-src/IsarOverview/Isar/document/Logic.tex Fri Aug 19 09:40:44 2005 +0200 +++ b/doc-src/IsarOverview/Isar/document/Logic.tex Fri Aug 19 10:50:05 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Logic}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsection{Logic \label{sec:Logic}% } @@ -19,8 +32,14 @@ We start with a really trivial toy proof to introduce the basic features of structured proofs.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}A\ {\isasymlongrightarrow}\ A{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\ {\isacharparenleft}rule\ impI{\isacharparenright}\isanewline \ \ \isamarkupfalse% @@ -29,7 +48,14 @@ \isacommand{show}\ {\isachardoublequote}A{\isachardoublequote}\ \isamarkupfalse% \isacommand{by}{\isacharparenleft}rule\ a{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -42,8 +68,14 @@ based on the goal and a predefined list of rules. \end{quote} Here \isa{impI} is applied automatically:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}A\ {\isasymlongrightarrow}\ A{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\isanewline \ \ \isamarkupfalse% @@ -52,7 +84,14 @@ \isacommand{show}\ A\ \isamarkupfalse% \isacommand{by}{\isacharparenleft}rule\ a{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Single-identifier formulae such as \isa{A} need not @@ -63,8 +102,14 @@ to perform. Proof ``.'' does just that (and a bit more). Thus naming of assumptions is often superfluous:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}A\ {\isasymlongrightarrow}\ A{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\isanewline \ \ \isamarkupfalse% @@ -73,15 +118,28 @@ \isacommand{show}\ {\isachardoublequote}A{\isachardoublequote}\ \isamarkupfalse% \isacommand{{\isachardot}}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% To hide proofs by assumption further, \isakeyword{by}\isa{{\isacharparenleft}method{\isacharparenright}} first applies \isa{method} and then tries to solve all remaining subgoals by assumption:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}A\ {\isasymlongrightarrow}\ A\ {\isasymand}\ A{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\isanewline \ \ \isamarkupfalse% @@ -90,7 +148,14 @@ \isacommand{show}\ {\isachardoublequote}A\ {\isasymand}\ A{\isachardoublequote}\ \isamarkupfalse% \isacommand{by}{\isacharparenleft}rule\ conjI{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Rule \isa{conjI} is of course \isa{{\isasymlbrakk}{\isacharquery}P{\isacharsemicolon}\ {\isacharquery}Q{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharquery}P\ {\isasymand}\ {\isacharquery}Q}. @@ -101,8 +166,14 @@ can be abbreviated to ``..'' if \emph{name} refers to one of the predefined introduction rules (or elimination rules, see below):% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}A\ {\isasymlongrightarrow}\ A\ {\isasymand}\ A{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\isanewline \ \ \isamarkupfalse% @@ -111,7 +182,14 @@ \isacommand{show}\ {\isachardoublequote}A\ {\isasymand}\ A{\isachardoublequote}\ \isamarkupfalse% \isacommand{{\isachardot}{\isachardot}}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -133,8 +211,14 @@ by hand, after its first (\emph{major}) premise has been eliminated via \isa{{\isacharbrackleft}OF\ AB{\isacharbrackright}}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}A\ {\isasymand}\ B\ {\isasymlongrightarrow}\ B\ {\isasymand}\ A{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\isanewline \ \ \isamarkupfalse% @@ -154,7 +238,14 @@ \ \ \isamarkupfalse% \isacommand{qed}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Note that the term \isa{{\isacharquery}thesis} always stands for the @@ -177,8 +268,14 @@ whose first premise is solved by the \emph{fact}. Thus the proof above is equivalent to the following one:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}A\ {\isasymand}\ B\ {\isasymlongrightarrow}\ B\ {\isasymand}\ A{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\isanewline \ \ \isamarkupfalse% @@ -196,7 +293,14 @@ \ \ \isamarkupfalse% \isacommand{qed}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% Now we come to a second important principle: @@ -207,8 +311,14 @@ The previous proposition can be referred to via the fact \isa{this}. This greatly reduces the need for explicit naming of propositions:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}A\ {\isasymand}\ B\ {\isasymlongrightarrow}\ B\ {\isasymand}\ A{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\isanewline \ \ \isamarkupfalse% @@ -226,7 +336,14 @@ \ \ \isamarkupfalse% \isacommand{qed}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Because of the frequency of \isakeyword{from}~\isa{this}, Isar provides two abbreviations: @@ -239,8 +356,14 @@ Here is an alternative proof that operates purely by forward reasoning:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}A\ {\isasymand}\ B\ {\isasymlongrightarrow}\ B\ {\isasymand}\ A{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\isanewline \ \ \isamarkupfalse% @@ -258,7 +381,14 @@ \isacommand{show}\ {\isachardoublequote}B\ {\isasymand}\ A{\isachardoublequote}\ \isamarkupfalse% \isacommand{{\isachardot}{\isachardot}}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent It is worth examining this text in detail because it @@ -306,8 +436,14 @@ UNIX-pipe model appears to break down and we need to name the different facts to refer to them. But this can be avoided:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}A\ {\isasymand}\ B\ {\isasymlongrightarrow}\ B\ {\isasymand}\ A{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\isanewline \ \ \isamarkupfalse% @@ -327,7 +463,14 @@ \isacommand{show}\ {\isachardoublequote}B\ {\isasymand}\ A{\isachardoublequote}\ \isamarkupfalse% \isacommand{{\isachardot}{\isachardot}}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent You can combine any number of facts \isa{A{\isadigit{1}}} \dots\ \isa{An} into a sequence by separating their proofs with @@ -343,8 +486,14 @@ rules. We conclude our exposition of propositional logic with an extended example --- which rules are used implicitly where?% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isasymnot}\ {\isacharparenleft}A\ {\isasymand}\ B{\isacharparenright}\ {\isasymlongrightarrow}\ {\isasymnot}\ A\ {\isasymor}\ {\isasymnot}\ B{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\isanewline \ \ \isamarkupfalse% @@ -395,7 +544,14 @@ \ \ \isamarkupfalse% \isacommand{qed}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -422,8 +578,14 @@ So far our examples have been a bit unnatural: normally we want to prove rules expressed with \isa{{\isasymLongrightarrow}}, not \isa{{\isasymlongrightarrow}}. Here is an example:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}A\ {\isasymand}\ B\ {\isasymLongrightarrow}\ B\ {\isasymand}\ A{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\isanewline \ \ \isamarkupfalse% @@ -437,7 +599,14 @@ \isacommand{thus}\ {\isachardoublequote}A{\isachardoublequote}\ \isamarkupfalse% \isacommand{{\isachardot}{\isachardot}}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent The \isakeyword{proof} always works on the conclusion, @@ -464,9 +633,15 @@ devices to avoid repeating (possibly large) formulae. A very general method is pattern matching:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}large{\isacharunderscore}A\ {\isasymand}\ large{\isacharunderscore}B\ {\isasymLongrightarrow}\ large{\isacharunderscore}B\ {\isasymand}\ large{\isacharunderscore}A{\isachardoublequote}\isanewline \ \ \ \ \ \ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequote}{\isacharquery}AB\ {\isasymLongrightarrow}\ {\isacharquery}B\ {\isasymand}\ {\isacharquery}A{\isachardoublequote}{\isacharparenright}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\isanewline \ \ \isamarkupfalse% @@ -480,7 +655,14 @@ \isacommand{thus}\ {\isachardoublequote}{\isacharquery}A{\isachardoublequote}\ \isamarkupfalse% \isacommand{{\isachardot}{\isachardot}}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Any formula may be followed by @@ -493,9 +675,15 @@ \isakeyword{assumes} and \isakeyword{shows} elements which allow direct naming of assumptions:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ \isakeyword{assumes}\ AB{\isacharcolon}\ {\isachardoublequote}large{\isacharunderscore}A\ {\isasymand}\ large{\isacharunderscore}B{\isachardoublequote}\isanewline \ \ \isakeyword{shows}\ {\isachardoublequote}large{\isacharunderscore}B\ {\isasymand}\ large{\isacharunderscore}A{\isachardoublequote}\ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequote}{\isacharquery}B\ {\isasymand}\ {\isacharquery}A{\isachardoublequote}{\isacharparenright}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\isanewline \ \ \isamarkupfalse% @@ -509,7 +697,14 @@ \isacommand{show}\ {\isachardoublequote}{\isacharquery}A{\isachardoublequote}\ \isamarkupfalse% \isacommand{{\isachardot}{\isachardot}}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Note the difference between \isa{{\isacharquery}AB}, a term, and @@ -519,9 +714,15 @@ don't have to perform it twice, as above. Here is a slick way to achieve this:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ \isakeyword{assumes}\ AB{\isacharcolon}\ {\isachardoublequote}large{\isacharunderscore}A\ {\isasymand}\ large{\isacharunderscore}B{\isachardoublequote}\isanewline \ \ \isakeyword{shows}\ {\isachardoublequote}large{\isacharunderscore}B\ {\isasymand}\ large{\isacharunderscore}A{\isachardoublequote}\ {\isacharparenleft}\isakeyword{is}\ {\isachardoublequote}{\isacharquery}B\ {\isasymand}\ {\isacharquery}A{\isachardoublequote}{\isacharparenright}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{using}\ AB\isanewline \isamarkupfalse% @@ -531,7 +732,14 @@ \isacommand{show}\ {\isacharquery}thesis\ \isamarkupfalse% \isacommand{{\isachardot}{\isachardot}}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Command \isakeyword{using} can appear before a proof @@ -550,8 +758,14 @@ trigger $\lor$-introduction, requiring us to prove \isa{A}. A simple ``\isa{{\isacharminus}}'' prevents this \emph{faux pas}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ \isakeyword{assumes}\ AB{\isacharcolon}\ {\isachardoublequote}A\ {\isasymor}\ B{\isachardoublequote}\ \isakeyword{shows}\ {\isachardoublequote}B\ {\isasymor}\ A{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\ {\isacharminus}\isanewline \ \ \isamarkupfalse% @@ -572,7 +786,14 @@ \ \ \isamarkupfalse% \isacommand{qed}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \isamarkupsubsection{Predicate calculus% } @@ -586,8 +807,14 @@ \isa{{\isasymLongrightarrow}}. Here is a sample proof, annotated with the rules that are applied implicitly:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ \isakeyword{assumes}\ P{\isacharcolon}\ {\isachardoublequote}{\isasymforall}x{\isachardot}\ P\ x{\isachardoublequote}\ \isakeyword{shows}\ {\isachardoublequote}{\isasymforall}x{\isachardot}\ P{\isacharparenleft}f\ x{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ % \isamarkupcmt{\isa{allI}: \isa{{\isacharparenleft}{\isasymAnd}x{\isachardot}\ {\isacharquery}P\ x{\isacharparenright}\ {\isasymLongrightarrow}\ {\isasymforall}x{\isachardot}\ {\isacharquery}P\ x}% @@ -603,7 +830,14 @@ } \isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Note that in the proof we have chosen to call the bound @@ -612,8 +846,14 @@ Next we look at \isa{{\isasymexists}} which is a bit more tricky.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ \isakeyword{assumes}\ Pf{\isacharcolon}\ {\isachardoublequote}{\isasymexists}x{\isachardot}\ P{\isacharparenleft}f\ x{\isacharparenright}{\isachardoublequote}\ \isakeyword{shows}\ {\isachardoublequote}{\isasymexists}y{\isachardot}\ P\ y{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\ {\isacharminus}\isanewline \ \ \isamarkupfalse% @@ -637,7 +877,14 @@ \ \ \isamarkupfalse% \isacommand{qed}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Explicit $\exists$-elimination as seen above can become @@ -645,8 +892,14 @@ \isakeyword{obtain} provides a more appealing form of generalised existence reasoning:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ \isakeyword{assumes}\ Pf{\isacharcolon}\ {\isachardoublequote}{\isasymexists}x{\isachardot}\ P{\isacharparenleft}f\ x{\isacharparenright}{\isachardoublequote}\ \isakeyword{shows}\ {\isachardoublequote}{\isasymexists}y{\isachardot}\ P\ y{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\ {\isacharminus}\isanewline \ \ \isamarkupfalse% @@ -657,7 +910,14 @@ \isacommand{thus}\ {\isachardoublequote}{\isasymexists}y{\isachardot}\ P\ y{\isachardoublequote}\ \isamarkupfalse% \isacommand{{\isachardot}{\isachardot}}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Note how the proof text follows the usual mathematical style @@ -669,8 +929,14 @@ Here is a proof of a well known tautology. Which rule is used where?% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ \isakeyword{assumes}\ ex{\isacharcolon}\ {\isachardoublequote}{\isasymexists}x{\isachardot}\ {\isasymforall}y{\isachardot}\ P\ x\ y{\isachardoublequote}\ \isakeyword{shows}\ {\isachardoublequote}{\isasymforall}y{\isachardot}\ {\isasymexists}x{\isachardot}\ P\ x\ y{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\isanewline \ \ \isamarkupfalse% @@ -686,7 +952,14 @@ \isacommand{thus}\ {\isachardoublequote}{\isasymexists}x{\isachardot}\ P\ x\ y{\isachardoublequote}\ \isamarkupfalse% \isacommand{{\isachardot}{\isachardot}}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \isamarkupsubsection{Making bigger steps% } @@ -698,8 +971,14 @@ Cantor's theorem that there is no surjective function from a set to its powerset:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{theorem}\ {\isachardoublequote}{\isasymexists}S{\isachardot}\ S\ {\isasymnotin}\ range\ {\isacharparenleft}f\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ set{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\isanewline \ \ \isamarkupfalse% @@ -737,7 +1016,14 @@ \ \ \isamarkupfalse% \isacommand{qed}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -758,8 +1044,14 @@ consumption, here is a more detailed version; the beginning up to \isakeyword{obtain} stays unchanged.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{theorem}\ {\isachardoublequote}{\isasymexists}S{\isachardot}\ S\ {\isasymnotin}\ range\ {\isacharparenleft}f\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ set{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\isanewline \ \ \isamarkupfalse% @@ -807,7 +1099,14 @@ \ \ \isamarkupfalse% \isacommand{qed}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Method \isa{contradiction} succeeds if both $P$ and @@ -817,10 +1116,23 @@ search. Depth-first search would diverge, but best-first search successfully navigates through the large search space:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{theorem}\ {\isachardoublequote}{\isasymexists}S{\isachardot}\ S\ {\isasymnotin}\ range\ {\isacharparenleft}f\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ set{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}\ best\isamarkupfalse% +\isacommand{by}\ best% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \isamarkupsubsection{Raw proof blocks% } @@ -832,9 +1144,27 @@ tendency to use the default introduction and elimination rules to decompose goals and facts. This can lead to very tedious proofs:% \end{isamarkuptext}% -\isamarkuptrue% +% +\isadelimML +% +\endisadelimML +% +\isatagML +% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isasymforall}x\ y{\isachardot}\ A\ x\ y\ {\isasymand}\ B\ x\ y\ {\isasymlongrightarrow}\ C\ x\ y{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\isanewline \ \ \isamarkupfalse% @@ -857,7 +1187,14 @@ \ \ \isamarkupfalse% \isacommand{qed}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Since we are only interested in the decomposition and not the @@ -868,8 +1205,14 @@ Luckily we can avoid this step by step decomposition very easily:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isasymforall}x\ y{\isachardot}\ A\ x\ y\ {\isasymand}\ B\ x\ y\ {\isasymlongrightarrow}\ C\ x\ y{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\ {\isacharminus}\isanewline \ \ \isamarkupfalse% @@ -888,15 +1231,28 @@ \isacommand{thus}\ {\isacharquery}thesis\ \isamarkupfalse% \isacommand{by}\ blast\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent This can be simplified further by \emph{raw proof blocks}, i.e.\ proofs enclosed in braces:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isasymforall}x\ y{\isachardot}\ A\ x\ y\ {\isasymand}\ B\ x\ y\ {\isasymlongrightarrow}\ C\ x\ y{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\ {\isacharminus}\isanewline \ \ \isamarkupfalse% @@ -911,7 +1267,14 @@ \isacommand{thus}\ {\isacharquery}thesis\ \isamarkupfalse% \isacommand{by}\ blast\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent The result of the raw proof block is the same theorem @@ -943,14 +1306,18 @@ that each case $P_i$ implies the goal. Taken together, this proves the goal. The corresponding Isar proof pattern (for $n = 3$) is very handy:% \end{isamarkuptext}% -\isamarkuptrue% % \renewcommand{\isamarkupcmt}[1]{#1} +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\ {\isacharminus}\isanewline \ \ \isamarkupfalse% -\isacommand{have}\ {\isachardoublequote}P\isactrlisub {\isadigit{1}}\ {\isasymor}\ P\isactrlisub {\isadigit{2}}\ {\isasymor}\ P\isactrlisub {\isadigit{3}}{\isachardoublequote}\ \isamarkupfalse% -\ % +\isacommand{have}\ {\isachardoublequote}P\isactrlisub {\isadigit{1}}\ {\isasymor}\ P\isactrlisub {\isadigit{2}}\ {\isasymor}\ P\isactrlisub {\isadigit{3}}{\isachardoublequote}\ \ % \isamarkupcmt{\dots% } \isanewline @@ -964,8 +1331,7 @@ } \isanewline \ \ \ \ \isamarkupfalse% -\isacommand{have}\ {\isacharquery}thesis\ \isamarkupfalse% -\ % +\isacommand{have}\ {\isacharquery}thesis\ \ % \isamarkupcmt{\dots% } \ \isamarkupfalse% @@ -980,8 +1346,7 @@ } \isanewline \ \ \ \ \isamarkupfalse% -\isacommand{have}\ {\isacharquery}thesis\ \isamarkupfalse% -\ % +\isacommand{have}\ {\isacharquery}thesis\ \ % \isamarkupcmt{\dots% } \ \isamarkupfalse% @@ -996,8 +1361,7 @@ } \isanewline \ \ \ \ \isamarkupfalse% -\isacommand{have}\ {\isacharquery}thesis\ \isamarkupfalse% -\ % +\isacommand{have}\ {\isacharquery}thesis\ \ % \isamarkupcmt{\dots% } \ \isamarkupfalse% @@ -1007,9 +1371,16 @@ \isacommand{show}\ {\isacharquery}thesis\ \isamarkupfalse% \isacommand{by}\ blast\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof % \renewcommand{\isamarkupcmt}[1]{{\isastylecmt--- #1}} +\isamarkuptrue% % \isamarkupsubsection{Further refinements% } @@ -1065,32 +1436,69 @@ to make the theorem less readable. The situation can be improved a little by combining the type constraint with an outer \isa{{\isasymAnd}}:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{theorem}\ {\isachardoublequote}{\isasymAnd}f\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ set{\isachardot}\ {\isasymexists}S{\isachardot}\ S\ {\isasymnotin}\ range\ f{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{theorem}\ {\isachardoublequote}{\isasymAnd}f\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ set{\isachardot}\ {\isasymexists}S{\isachardot}\ S\ {\isasymnotin}\ range\ f{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% % \begin{isamarkuptext}% \noindent However, now \isa{f} is bound and we need a \isakeyword{fix}~\isa{f} in the proof before we can refer to \isa{f}. This is avoided by \isakeyword{fixes}:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{theorem}\ \isakeyword{fixes}\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ set{\isachardoublequote}\ \isakeyword{shows}\ {\isachardoublequote}{\isasymexists}S{\isachardot}\ S\ {\isasymnotin}\ range\ f{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{theorem}\ \isakeyword{fixes}\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ set{\isachardoublequote}\ \isakeyword{shows}\ {\isachardoublequote}{\isasymexists}S{\isachardot}\ S\ {\isasymnotin}\ range\ f{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% % \begin{isamarkuptext}% \noindent Even better, \isakeyword{fixes} allows to introduce concrete syntax locally:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ comm{\isacharunderscore}mono{\isacharcolon}\isanewline \ \ \isakeyword{fixes}\ r\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ bool{\isachardoublequote}\ {\isacharparenleft}\isakeyword{infix}\ {\isachardoublequote}{\isachargreater}{\isachardoublequote}\ {\isadigit{6}}{\isadigit{0}}{\isacharparenright}\ \isakeyword{and}\isanewline \ \ \ \ \ \ \ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isacharplus}{\isacharplus}{\isachardoublequote}\ {\isadigit{7}}{\isadigit{0}}{\isacharparenright}\isanewline \ \ \isakeyword{assumes}\ comm{\isacharcolon}\ {\isachardoublequote}{\isasymAnd}x\ y{\isacharcolon}{\isacharcolon}{\isacharprime}a{\isachardot}\ x\ {\isacharplus}{\isacharplus}\ y\ {\isacharequal}\ y\ {\isacharplus}{\isacharplus}\ x{\isachardoublequote}\ \isakeyword{and}\isanewline \ \ \ \ \ \ \ \ \ \ mono{\isacharcolon}\ {\isachardoublequote}{\isasymAnd}x\ y\ z{\isacharcolon}{\isacharcolon}{\isacharprime}a{\isachardot}\ x\ {\isachargreater}\ y\ {\isasymLongrightarrow}\ x\ {\isacharplus}{\isacharplus}\ z\ {\isachargreater}\ y\ {\isacharplus}{\isacharplus}\ z{\isachardoublequote}\isanewline \ \ \isakeyword{shows}\ {\isachardoublequote}x\ {\isachargreater}\ y\ {\isasymLongrightarrow}\ z\ {\isacharplus}{\isacharplus}\ x\ {\isachargreater}\ z\ {\isacharplus}{\isacharplus}\ y{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}{\isacharparenleft}simp\ add{\isacharcolon}\ comm\ mono{\isacharparenright}\isamarkupfalse% +\isacommand{by}{\isacharparenleft}simp\ add{\isacharcolon}\ comm\ mono{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent The concrete syntax is dropped at the end of the proof and the @@ -1111,15 +1519,27 @@ The \isakeyword{obtain} construct can introduce multiple witnesses and propositions as in the following proof fragment:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ \isakeyword{assumes}\ A{\isacharcolon}\ {\isachardoublequote}{\isasymexists}x\ y{\isachardot}\ P\ x\ y\ {\isasymand}\ Q\ x\ y{\isachardoublequote}\ \isakeyword{shows}\ {\isachardoublequote}R{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\ {\isacharminus}\isanewline \ \ \isamarkupfalse% \isacommand{from}\ A\ \isamarkupfalse% \isacommand{obtain}\ x\ y\ \isakeyword{where}\ P{\isacharcolon}\ {\isachardoublequote}P\ x\ y{\isachardoublequote}\ \isakeyword{and}\ Q{\isacharcolon}\ {\isachardoublequote}Q\ x\ y{\isachardoublequote}\ \ \isamarkupfalse% -\isacommand{by}\ blast\isamarkupfalse% -\isamarkupfalse% +\isacommand{by}\ blast% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% Remember also that one does not even need to start with a formula @@ -1136,8 +1556,14 @@ \cite{LNCS2283}) may appear in the leaves of the proof tree, although this is best avoided. Here is a contrived example:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}A\ {\isasymlongrightarrow}\ {\isacharparenleft}A\ {\isasymlongrightarrow}\ B{\isacharparenright}\ {\isasymlongrightarrow}\ B{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\isanewline \ \ \isamarkupfalse% @@ -1155,7 +1581,14 @@ \ \ \ \ \isamarkupfalse% \isacommand{done}\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent You may need to resort to this technique if an @@ -1165,8 +1598,19 @@ in a top-down manner: parts of the proof can be left in script form while the outer structure is already expressed in Isar.% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r 19ea4a0f4ec9 -r e6a82d1a1829 doc-src/IsarOverview/Isar/document/isabelle.sty --- a/doc-src/IsarOverview/Isar/document/isabelle.sty Fri Aug 19 09:40:44 2005 +0200 +++ b/doc-src/IsarOverview/Isar/document/isabelle.sty Fri Aug 19 10:50:05 2005 +0200 @@ -22,10 +22,10 @@ \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{\isactrlsub}[1]{\emph{${}\sb{\mbox{\isastylescript{#1}}}$}} +\newcommand{\isactrlsup}[1]{\emph{${}\sp{\mbox{\isastylescript{#1}}}$}} +\newcommand{\isactrlisub}[1]{\emph{${}\sb{\mbox{\isastylescript{#1}}}$}} +\newcommand{\isactrlisup}[1]{\emph{${}\sp{\mbox{\isastylescript{#1}}}$}} \newcommand{\isactrlbsub}{\emph\bgroup\begin{math}{}\sb\bgroup\mbox\bgroup\isastylescript} \newcommand{\isactrlesub}{\egroup\egroup\end{math}\egroup} \newcommand{\isactrlbsup}{\emph\bgroup\begin{math}{}\sp\bgroup\mbox\bgroup\isastylescript} @@ -149,6 +149,7 @@ \renewcommand{\isachargreater}{\isamath{>}}% \renewcommand{\isacharat}{\isamath{@}}% \renewcommand{\isacharbrackleft}{\isamath{[}}% +\renewcommand{\isacharbackslash}{\isamath{\backslash}}% \renewcommand{\isacharbrackright}{\isamath{]}}% \renewcommand{\isacharunderscore}{\mbox{-}}% \renewcommand{\isacharbraceleft}{\isamath{\{}}% @@ -163,3 +164,28 @@ \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 19ea4a0f4ec9 -r e6a82d1a1829 doc-src/IsarOverview/Isar/document/isabellesym.sty --- a/doc-src/IsarOverview/Isar/document/isabellesym.sty Fri Aug 19 09:40:44 2005 +0200 +++ b/doc-src/IsarOverview/Isar/document/isabellesym.sty Fri Aug 19 10:50:05 2005 +0200 @@ -359,3 +359,4 @@ \newcommand{\isasymcedilla}{\isatext{\c\relax}} \newcommand{\isasymhungarumlaut}{\isatext{\H\relax}} \newcommand{\isasymspacespace}{\isamath{~~}} +\newcommand{\isasymmodule}{\isamath{\langle}\isakeyword{module}\isamath{\rangle}} diff -r 19ea4a0f4ec9 -r e6a82d1a1829 doc-src/LaTeXsugar/Sugar/document/Sugar.tex --- a/doc-src/LaTeXsugar/Sugar/document/Sugar.tex Fri Aug 19 09:40:44 2005 +0200 +++ b/doc-src/LaTeXsugar/Sugar/document/Sugar.tex Fri Aug 19 10:50:05 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Sugar}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsection{Introduction% } @@ -150,8 +163,20 @@ internal index. This can be avoided by turning the last digit into a subscript: write \verb!x\<^isub>1! and obtain the much nicer \isa{x\isactrlisub {\isadigit{1}}}.% \end{isamarkuptext}% +% +\isadelimML +% +\endisadelimML +% +\isatagML +% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML \isamarkuptrue% -\isamarkupfalse% % \isamarkupsubsection{Variable names% } @@ -329,12 +354,18 @@ \ref{fig:proof}. This was achieved with the \isakeyword{text\_raw} command:% \end{isamarkuptext}% -\isamarkuptrue% % \begin{figure} \begin{center}\begin{minipage}{0.6\textwidth} \isastyle\isamarkuptrue +\isamarkupfalse% \isacommand{lemma}\ True\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{proof}\ {\isacharminus}\isanewline \ \ % @@ -345,11 +376,18 @@ \isacommand{show}\ True\ \isamarkupfalse% \isacommand{by}\ force\isanewline \isamarkupfalse% -\isacommand{qed}\isamarkupfalse% +\isacommand{qed}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof % \end{minipage}\end{center} \caption{Example proof in a figure.}\label{fig:proof} \end{figure} +\isamarkuptrue% % \begin{isamarkuptext}% \begin{quote} @@ -420,19 +458,22 @@ Likewise, \verb!concl! may be used as a style to show just the conclusion of a proposition. For example, take \verb!hd_Cons_tl!: \begin{center} - \isa{xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymLongrightarrow}\ hd\ xs{\isasymcdot}tl\ xs\ {\isacharequal}\ xs} + \isa{{\isacharparenleft}xs{\isasymColon}{\isacharprime}a\ list{\isacharparenright}\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymLongrightarrow}\ hd\ xs{\isasymcdot}tl\ xs\ {\isacharequal}\ xs} \end{center} To print just the conclusion, \begin{center} - \isa{hd\ xs{\isasymcdot}tl\ xs\ {\isacharequal}\ xs} + \isa{hd\ {\isacharparenleft}xs{\isasymColon}{\isacharprime}a\ list{\isacharparenright}{\isasymcdot}tl\ xs\ {\isacharequal}\ xs} \end{center} type \begin{quote} \verb!\begin{center}!\\ - \verb!@!\verb!{thm_style concl hd_Cons_tl}!\\ + \verb!@!\verb!{thm_style [show_types] concl hd_Cons_tl}!\\ \verb!\end{center}! \end{quote} + Be aware that any options must be placed \emph{before} + the name of the style, as in this example. + Further use cases can be found in \S\ref{sec:yourself}. If you are not afraid of ML, you may also define your own styles. @@ -440,8 +481,20 @@ \verb!Proof.context -> term -> term!. Have a look at the following example:% \end{isamarkuptext}% +% +\isadelimML +% +\endisadelimML +% +\isatagML +% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML \isamarkuptrue% -\isamarkupfalse% % \begin{isamarkuptext}% \begin{quote} @@ -464,7 +517,7 @@ style has some object-logic specific behaviour for example. The mapping from identifier name to the style function - is done by the \isa{TermStyle{\isachardot}add{\isacharunderscore}style} expression which expects the desired + is done by the \verb|TermStyle.add_style| expression which expects the desired style name and the style function as arguments. After this \verb!setup!, @@ -472,8 +525,19 @@ antiquoations like \verb!@!\verb!{thm_style my_concl hd_Cons_tl}! yielding \isa{hd\ xs{\isasymcdot}tl\ xs\ {\isacharequal}\ xs}.% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r 19ea4a0f4ec9 -r e6a82d1a1829 doc-src/LaTeXsugar/Sugar/document/isabelle.sty --- a/doc-src/LaTeXsugar/Sugar/document/isabelle.sty Fri Aug 19 09:40:44 2005 +0200 +++ b/doc-src/LaTeXsugar/Sugar/document/isabelle.sty Fri Aug 19 10:50:05 2005 +0200 @@ -22,10 +22,10 @@ \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{\isactrlsub}[1]{\emph{${}\sb{\mbox{\isastylescript{#1}}}$}} +\newcommand{\isactrlsup}[1]{\emph{${}\sp{\mbox{\isastylescript{#1}}}$}} +\newcommand{\isactrlisub}[1]{\emph{${}\sb{\mbox{\isastylescript{#1}}}$}} +\newcommand{\isactrlisup}[1]{\emph{${}\sp{\mbox{\isastylescript{#1}}}$}} \newcommand{\isactrlbsub}{\emph\bgroup\begin{math}{}\sb\bgroup\mbox\bgroup\isastylescript} \newcommand{\isactrlesub}{\egroup\egroup\end{math}\egroup} \newcommand{\isactrlbsup}{\emph\bgroup\begin{math}{}\sp\bgroup\mbox\bgroup\isastylescript} @@ -149,6 +149,7 @@ \renewcommand{\isachargreater}{\isamath{>}}% \renewcommand{\isacharat}{\isamath{@}}% \renewcommand{\isacharbrackleft}{\isamath{[}}% +\renewcommand{\isacharbackslash}{\isamath{\backslash}}% \renewcommand{\isacharbrackright}{\isamath{]}}% \renewcommand{\isacharunderscore}{\mbox{-}}% \renewcommand{\isacharbraceleft}{\isamath{\{}}% @@ -163,3 +164,28 @@ \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 19ea4a0f4ec9 -r e6a82d1a1829 doc-src/LaTeXsugar/Sugar/document/isabellesym.sty --- a/doc-src/LaTeXsugar/Sugar/document/isabellesym.sty Fri Aug 19 09:40:44 2005 +0200 +++ b/doc-src/LaTeXsugar/Sugar/document/isabellesym.sty Fri Aug 19 10:50:05 2005 +0200 @@ -359,3 +359,4 @@ \newcommand{\isasymcedilla}{\isatext{\c\relax}} \newcommand{\isasymhungarumlaut}{\isatext{\H\relax}} \newcommand{\isasymspacespace}{\isamath{~~}} +\newcommand{\isasymmodule}{\isamath{\langle}\isakeyword{module}\isamath{\rangle}}