# HG changeset patch # User wenzelm # Date 1124192543 -7200 # Node ID 05fc32a23b8b2443261e3c62b33a4407c42e3664 # Parent eacce1cd716a0f4ee75a73703863526d6110706b updated; diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Advanced/document/Partial.tex --- a/doc-src/TutorialI/Advanced/document/Partial.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Advanced/document/Partial.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Partial}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Throughout this tutorial, we have emphasized @@ -28,10 +41,10 @@ non-exhaustive pattern matching: the definition of \isa{last} in \S\ref{sec:recdef-examples}. The same is allowed for \isacommand{primrec}% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ hd\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{primrec}\ {\isachardoublequote}hd\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ x{\isachardoublequote}\isamarkupfalse% +\isacommand{primrec}\ {\isachardoublequote}hd\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ x{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -39,9 +52,9 @@ Even ordinary definitions allow underdefinedness, this time by means of preconditions:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{constdefs}\ minus\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequote}\isanewline -{\isachardoublequote}n\ {\isasymle}\ m\ {\isasymLongrightarrow}\ minus\ m\ n\ {\isasymequiv}\ m\ {\isacharminus}\ n{\isachardoublequote}\isamarkupfalse% +{\isachardoublequote}n\ {\isasymle}\ m\ {\isasymLongrightarrow}\ minus\ m\ n\ {\isasymequiv}\ m\ {\isacharminus}\ n{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% The rest of this section is devoted to the question of how to define @@ -71,12 +84,12 @@ As a simple example we define division on \isa{nat}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ divi\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat\ {\isasymtimes}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{recdef}\ divi\ {\isachardoublequote}measure{\isacharparenleft}{\isasymlambda}{\isacharparenleft}m{\isacharcomma}n{\isacharparenright}{\isachardot}\ m{\isacharparenright}{\isachardoublequote}\isanewline \ \ {\isachardoublequote}divi{\isacharparenleft}m{\isacharcomma}{\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ arbitrary{\isachardoublequote}\isanewline -\ \ {\isachardoublequote}divi{\isacharparenleft}m{\isacharcomma}n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ m\ {\isacharless}\ n\ then\ {\isadigit{0}}\ else\ divi{\isacharparenleft}m{\isacharminus}n{\isacharcomma}n{\isacharparenright}{\isacharplus}{\isadigit{1}}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}divi{\isacharparenleft}m{\isacharcomma}n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ m\ {\isacharless}\ n\ then\ {\isadigit{0}}\ else\ divi{\isacharparenleft}m{\isacharminus}n{\isacharcomma}n{\isacharparenright}{\isacharplus}{\isadigit{1}}{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent Of course we could also have defined @@ -99,22 +112,22 @@ The snag is that it may not terminate if \isa{f} has non-trivial cycles. Phrased differently, the relation% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{constdefs}\ step{\isadigit{1}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}set{\isachardoublequote}\isanewline -\ \ {\isachardoublequote}step{\isadigit{1}}\ f\ {\isasymequiv}\ {\isacharbraceleft}{\isacharparenleft}y{\isacharcomma}x{\isacharparenright}{\isachardot}\ y\ {\isacharequal}\ f\ x\ {\isasymand}\ y\ {\isasymnoteq}\ x{\isacharbraceright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}step{\isadigit{1}}\ f\ {\isasymequiv}\ {\isacharbraceleft}{\isacharparenleft}y{\isacharcomma}x{\isacharparenright}{\isachardot}\ y\ {\isacharequal}\ f\ x\ {\isasymand}\ y\ {\isasymnoteq}\ x{\isacharbraceright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent must be well-founded. Thus we make the following definition:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ find\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isasymtimes}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{recdef}\ find\ {\isachardoublequote}same{\isacharunderscore}fst\ {\isacharparenleft}{\isasymlambda}f{\isachardot}\ wf{\isacharparenleft}step{\isadigit{1}}\ f{\isacharparenright}{\isacharparenright}\ step{\isadigit{1}}{\isachardoublequote}\isanewline \ \ {\isachardoublequote}find{\isacharparenleft}f{\isacharcomma}x{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ wf{\isacharparenleft}step{\isadigit{1}}\ f{\isacharparenright}\isanewline \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ then\ if\ f\ x\ {\isacharequal}\ x\ then\ x\ else\ find{\isacharparenleft}f{\isacharcomma}\ f\ x{\isacharparenright}\isanewline \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ else\ arbitrary{\isacharparenright}{\isachardoublequote}\isanewline -{\isacharparenleft}\isakeyword{hints}\ recdef{\isacharunderscore}simp{\isacharcolon}\ step{\isadigit{1}}{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% +{\isacharparenleft}\isakeyword{hints}\ recdef{\isacharunderscore}simp{\isacharcolon}\ step{\isadigit{1}}{\isacharunderscore}def{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -146,30 +159,56 @@ Normally you will then derive the following conditional variant from the recursion equation:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\isanewline \ \ {\isachardoublequote}wf{\isacharparenleft}step{\isadigit{1}}\ f{\isacharparenright}\ {\isasymLongrightarrow}\ find{\isacharparenleft}f{\isacharcomma}x{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ f\ x\ {\isacharequal}\ x\ then\ x\ else\ find{\isacharparenleft}f{\isacharcomma}\ f\ x{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +\isacommand{by}\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Then you should disable the original recursion equation:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{declare}\ find{\isachardot}simps{\isacharbrackleft}simp\ del{\isacharbrackright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{declare}\ find{\isachardot}simps{\isacharbrackleft}simp\ del{\isacharbrackright}\isamarkuptrue% % \begin{isamarkuptext}% Reasoning about such underdefined functions is like that for other recursive functions. Here is a simple example of recursion induction:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}wf{\isacharparenleft}step{\isadigit{1}}\ f{\isacharparenright}\ {\isasymlongrightarrow}\ f{\isacharparenleft}find{\isacharparenleft}f{\isacharcomma}x{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ find{\isacharparenleft}f{\isacharcomma}x{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ f\ x\ rule{\isacharcolon}\ find{\isachardot}induct{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}\ simp\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \isamarkupsubsubsection{The {\tt\slshape while} Combinator% } @@ -193,10 +232,10 @@ In general, \isa{s} will be a tuple or record. As an example consider the following definition of function \isa{find}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{constdefs}\ find{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\isanewline \ \ {\isachardoublequote}find{\isadigit{2}}\ f\ x\ {\isasymequiv}\isanewline -\ \ \ fst{\isacharparenleft}while\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}x{\isacharcomma}x{\isacharprime}{\isacharparenright}{\isachardot}\ x{\isacharprime}\ {\isasymnoteq}\ x{\isacharparenright}\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}x{\isacharcomma}x{\isacharprime}{\isacharparenright}{\isachardot}\ {\isacharparenleft}x{\isacharprime}{\isacharcomma}f\ x{\isacharprime}{\isacharparenright}{\isacharparenright}\ {\isacharparenleft}x{\isacharcomma}f\ x{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ \ fst{\isacharparenleft}while\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}x{\isacharcomma}x{\isacharprime}{\isacharparenright}{\isachardot}\ x{\isacharprime}\ {\isasymnoteq}\ x{\isacharparenright}\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}x{\isacharcomma}x{\isacharprime}{\isacharparenright}{\isachardot}\ {\isacharparenleft}x{\isacharprime}{\isacharcomma}f\ x{\isacharprime}{\isacharparenright}{\isacharparenright}\ {\isacharparenleft}x{\isacharcomma}f\ x{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -226,10 +265,16 @@ Only the final premise of \isa{while{\isacharunderscore}rule} is left unproved by \isa{auto} but falls to \isa{simp}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ lem{\isacharcolon}\ {\isachardoublequote}wf{\isacharparenleft}step{\isadigit{1}}\ f{\isacharparenright}\ {\isasymLongrightarrow}\isanewline \ \ {\isasymexists}y{\isachardot}\ while\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}x{\isacharcomma}x{\isacharprime}{\isacharparenright}{\isachardot}\ x{\isacharprime}\ {\isasymnoteq}\ x{\isacharparenright}\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}x{\isacharcomma}x{\isacharprime}{\isacharparenright}{\isachardot}\ {\isacharparenleft}x{\isacharprime}{\isacharcomma}f\ x{\isacharprime}{\isacharparenright}{\isacharparenright}\ {\isacharparenleft}x{\isacharcomma}f\ x{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}y{\isacharcomma}y{\isacharparenright}\ {\isasymand}\isanewline \ \ \ \ \ \ \ f\ y\ {\isacharequal}\ y{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}rule{\isacharunderscore}tac\ P\ {\isacharequal}\ {\isachardoublequote}{\isasymlambda}{\isacharparenleft}x{\isacharcomma}x{\isacharprime}{\isacharparenright}{\isachardot}\ x{\isacharprime}\ {\isacharequal}\ f\ x{\isachardoublequote}\ \isakeyword{and}\isanewline \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ r\ {\isacharequal}\ {\isachardoublequote}inv{\isacharunderscore}image\ {\isacharparenleft}step{\isadigit{1}}\ f{\isacharparenright}\ fst{\isachardoublequote}\ \isakeyword{in}\ while{\isacharunderscore}rule{\isacharparenright}\isanewline @@ -238,19 +283,39 @@ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}simp\ add{\isacharcolon}\ inv{\isacharunderscore}image{\isacharunderscore}def\ step{\isadigit{1}}{\isacharunderscore}def{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% The theorem itself is a simple consequence of this lemma:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{theorem}\ {\isachardoublequote}wf{\isacharparenleft}step{\isadigit{1}}\ f{\isacharparenright}\ {\isasymLongrightarrow}\ f{\isacharparenleft}find{\isadigit{2}}\ f\ x{\isacharparenright}\ {\isacharequal}\ find{\isadigit{2}}\ f\ x{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}drule{\isacharunderscore}tac\ x\ {\isacharequal}\ x\ \isakeyword{in}\ lem{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}{\isacharparenleft}auto\ simp\ add{\isacharcolon}\ find{\isadigit{2}}{\isacharunderscore}def{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% Let us conclude this section on partial functions by a @@ -267,8 +332,19 @@ Thus, if you are aiming for an efficiently executable definition of a partial function, you are likely to need \isa{while}.% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Advanced/document/WFrec.tex --- a/doc-src/TutorialI/Advanced/document/WFrec.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Advanced/document/WFrec.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{WFrec}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -11,13 +24,13 @@ general definitions. For example, termination of Ackermann's function can be shown by means of the \rmindex{lexicographic product} \isa{{\isacharless}{\isacharasterisk}lex{\isacharasterisk}{\isachargreater}}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ ack\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat{\isasymtimes}nat\ {\isasymRightarrow}\ nat{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{recdef}\ ack\ {\isachardoublequote}measure{\isacharparenleft}{\isasymlambda}m{\isachardot}\ m{\isacharparenright}\ {\isacharless}{\isacharasterisk}lex{\isacharasterisk}{\isachargreater}\ measure{\isacharparenleft}{\isasymlambda}n{\isachardot}\ n{\isacharparenright}{\isachardoublequote}\isanewline \ \ {\isachardoublequote}ack{\isacharparenleft}{\isadigit{0}}{\isacharcomma}n{\isacharparenright}\ \ \ \ \ \ \ \ \ {\isacharequal}\ Suc\ n{\isachardoublequote}\isanewline \ \ {\isachardoublequote}ack{\isacharparenleft}Suc\ m{\isacharcomma}{\isadigit{0}}{\isacharparenright}\ \ \ \ \ {\isacharequal}\ ack{\isacharparenleft}m{\isacharcomma}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequote}\isanewline -\ \ {\isachardoublequote}ack{\isacharparenleft}Suc\ m{\isacharcomma}Suc\ n{\isacharparenright}\ {\isacharequal}\ ack{\isacharparenleft}m{\isacharcomma}ack{\isacharparenleft}Suc\ m{\isacharcomma}n{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}ack{\isacharparenleft}Suc\ m{\isacharcomma}Suc\ n{\isacharparenright}\ {\isacharequal}\ ack{\isacharparenleft}m{\isacharcomma}ack{\isacharparenleft}Suc\ m{\isacharcomma}n{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -46,7 +59,7 @@ on when defining Ackermann's function above. Of course the lexicographic product can also be iterated:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ contrived\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat\ {\isasymtimes}\ nat\ {\isasymtimes}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{recdef}\ contrived\isanewline @@ -54,7 +67,7 @@ {\isachardoublequote}contrived{\isacharparenleft}i{\isacharcomma}j{\isacharcomma}Suc\ k{\isacharparenright}\ {\isacharequal}\ contrived{\isacharparenleft}i{\isacharcomma}j{\isacharcomma}k{\isacharparenright}{\isachardoublequote}\isanewline {\isachardoublequote}contrived{\isacharparenleft}i{\isacharcomma}Suc\ j{\isacharcomma}{\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ contrived{\isacharparenleft}i{\isacharcomma}j{\isacharcomma}j{\isacharparenright}{\isachardoublequote}\isanewline {\isachardoublequote}contrived{\isacharparenleft}Suc\ i{\isacharcomma}{\isadigit{0}}{\isacharcomma}{\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ contrived{\isacharparenleft}i{\isacharcomma}i{\isacharcomma}i{\isacharparenright}{\isachardoublequote}\isanewline -{\isachardoublequote}contrived{\isacharparenleft}{\isadigit{0}}{\isacharcomma}{\isadigit{0}}{\isacharcomma}{\isadigit{0}}{\isacharparenright}\ \ \ \ \ {\isacharequal}\ {\isadigit{0}}{\isachardoublequote}\isamarkupfalse% +{\isachardoublequote}contrived{\isacharparenleft}{\isadigit{0}}{\isacharcomma}{\isadigit{0}}{\isacharcomma}{\isadigit{0}}{\isacharparenright}\ \ \ \ \ {\isacharequal}\ {\isadigit{0}}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% Lexicographic products of measure functions already go a long @@ -71,11 +84,11 @@ well-founded by cutting it off at a certain point. Here is an example of a recursive function that calls itself with increasing values up to ten:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat\ {\isasymRightarrow}\ nat{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{recdef}\ f\ {\isachardoublequote}{\isacharbraceleft}{\isacharparenleft}i{\isacharcomma}j{\isacharparenright}{\isachardot}\ j{\isacharless}i\ {\isasymand}\ i\ {\isasymle}\ {\isacharparenleft}{\isadigit{1}}{\isadigit{0}}{\isacharcolon}{\isacharcolon}nat{\isacharparenright}{\isacharbraceright}{\isachardoublequote}\isanewline -{\isachardoublequote}f\ i\ {\isacharequal}\ {\isacharparenleft}if\ {\isadigit{1}}{\isadigit{0}}\ {\isasymle}\ i\ then\ {\isadigit{0}}\ else\ i\ {\isacharasterisk}\ f{\isacharparenleft}Suc\ i{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +{\isachardoublequote}f\ i\ {\isacharequal}\ {\isacharparenleft}if\ {\isadigit{1}}{\isadigit{0}}\ {\isasymle}\ i\ then\ {\isadigit{0}}\ else\ i\ {\isacharasterisk}\ f{\isacharparenleft}Suc\ i{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -83,16 +96,22 @@ Isabelle rejects the definition. We should first have proved that our relation was well-founded:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ wf{\isacharunderscore}greater{\isacharcolon}\ {\isachardoublequote}wf\ {\isacharbraceleft}{\isacharparenleft}i{\isacharcomma}j{\isacharparenright}{\isachardot}\ j{\isacharless}i\ {\isasymand}\ i\ {\isasymle}\ {\isacharparenleft}N{\isacharcolon}{\isacharcolon}nat{\isacharparenright}{\isacharbraceright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{lemma}\ wf{\isacharunderscore}greater{\isacharcolon}\ {\isachardoublequote}wf\ {\isacharbraceleft}{\isacharparenleft}i{\isacharcomma}j{\isacharparenright}{\isachardot}\ j{\isacharless}i\ {\isasymand}\ i\ {\isasymle}\ {\isacharparenleft}N{\isacharcolon}{\isacharcolon}nat{\isacharparenright}{\isacharbraceright}{\isachardoublequote}\isamarkupfalse% % \begin{isamarkuptxt}% \noindent The proof is by showing that our relation is a subset of another well-founded relation: one given by a measure function.\index{*wf_subset (theorem)}% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\ {\isacharparenleft}rule\ wf{\isacharunderscore}subset\ {\isacharbrackleft}of\ {\isachardoublequote}measure\ {\isacharparenleft}{\isasymlambda}k{\isacharcolon}{\isacharcolon}nat{\isachardot}\ N{\isacharminus}k{\isacharparenright}{\isachardoublequote}{\isacharbrackright}{\isacharcomma}\ blast{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}\ {\isacharparenleft}rule\ wf{\isacharunderscore}subset\ {\isacharbrackleft}of\ {\isachardoublequote}measure\ {\isacharparenleft}{\isasymlambda}k{\isacharcolon}{\isacharcolon}nat{\isachardot}\ N{\isacharminus}k{\isacharparenright}{\isachardoublequote}{\isacharbrackright}{\isacharcomma}\ blast{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -103,8 +122,8 @@ The inclusion remains to be proved. After unfolding some definitions, we are left with simple arithmetic:% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\ {\isacharparenleft}clarify{\isacharcomma}\ simp\ add{\isacharcolon}\ measure{\isacharunderscore}def\ inv{\isacharunderscore}image{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}\ {\isacharparenleft}clarify{\isacharcomma}\ simp\ add{\isacharcolon}\ measure{\isacharunderscore}def\ inv{\isacharunderscore}image{\isacharunderscore}def{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -114,8 +133,15 @@ \noindent And that is dispatched automatically:% \end{isamarkuptxt}% +\isamarkupfalse% +\isacommand{by}\ arith% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{by}\ arith\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -123,9 +149,7 @@ Armed with this lemma, we use the \attrdx{recdef_wf} attribute to attach a crucial hint\cmmdx{hints} to our definition:% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% -{\isacharparenleft}\isakeyword{hints}\ recdef{\isacharunderscore}wf{\isacharcolon}\ wf{\isacharunderscore}greater{\isacharparenright}\isamarkupfalse% +{\isacharparenleft}\isakeyword{hints}\ recdef{\isacharunderscore}wf{\isacharcolon}\ wf{\isacharunderscore}greater{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -136,8 +160,19 @@ relation makes even more sense when it can be used in several function declarations.% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Advanced/document/simp.tex --- a/doc-src/TutorialI/Advanced/document/simp.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Advanced/document/simp.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{simp}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsection{Simplification% } @@ -216,8 +229,19 @@ \index{simplification rule|)} \index{simplification|)}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/CTL/document/Base.tex --- a/doc-src/TutorialI/CTL/document/Base.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/CTL/document/Base.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Base}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsection{Case Study: Verified Model Checking% } @@ -61,8 +74,8 @@ Abstracting from this concrete example, we assume there is a type of states:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{typedecl}\ state\isamarkupfalse% +\isamarkupfalse% +\isacommand{typedecl}\ state\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -75,30 +88,41 @@ reduces clutter. Similarly we declare an arbitrary but fixed transition system, i.e.\ a relation between states:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{consts}\ M\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}state\ {\isasymtimes}\ state{\isacharparenright}set{\isachardoublequote}\isamarkupfalse% +\isamarkupfalse% +\isacommand{consts}\ M\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}state\ {\isasymtimes}\ state{\isacharparenright}set{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent Again, we could have made \isa{M} a parameter of everything. Finally we introduce a type of atomic propositions% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{typedecl}\ atom\isamarkupfalse% +\isamarkupfalse% +\isacommand{typedecl}\ atom\isamarkuptrue% % \begin{isamarkuptext}% \noindent and a \emph{labelling function}% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{consts}\ L\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}state\ {\isasymRightarrow}\ atom\ set{\isachardoublequote}\isamarkupfalse% +\isamarkupfalse% +\isacommand{consts}\ L\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}state\ {\isasymRightarrow}\ atom\ set{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent telling us which atomic propositions are true in each state.% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/CTL/document/CTL.tex --- a/doc-src/TutorialI/CTL/document/CTL.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/CTL/document/CTL.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{CTL}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsubsection{Computation Tree Logic --- CTL% } @@ -15,8 +28,7 @@ We extend the datatype \isa{formula} by a new constructor% \end{isamarkuptext}% -\isamarkuptrue% -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ AF\ formula\isamarkupfalse% +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ AF\ formula\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -25,9 +37,9 @@ Formalizing the notion of an infinite path is easy in HOL: it is simply a function from \isa{nat} to \isa{state}.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{constdefs}\ Paths\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}state\ {\isasymRightarrow}\ {\isacharparenleft}nat\ {\isasymRightarrow}\ state{\isacharparenright}set{\isachardoublequote}\isanewline -\ \ \ \ \ \ \ \ \ {\isachardoublequote}Paths\ s\ {\isasymequiv}\ {\isacharbraceleft}p{\isachardot}\ s\ {\isacharequal}\ p\ {\isadigit{0}}\ {\isasymand}\ {\isacharparenleft}{\isasymforall}i{\isachardot}\ {\isacharparenleft}p\ i{\isacharcomma}\ p{\isacharparenleft}i{\isacharplus}{\isadigit{1}}{\isacharparenright}{\isacharparenright}\ {\isasymin}\ M{\isacharparenright}{\isacharbraceright}{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ \ \ \ \ {\isachardoublequote}Paths\ s\ {\isasymequiv}\ {\isacharbraceleft}p{\isachardot}\ s\ {\isacharequal}\ p\ {\isadigit{0}}\ {\isasymand}\ {\isacharparenleft}{\isasymforall}i{\isachardot}\ {\isacharparenleft}p\ i{\isacharcomma}\ p{\isacharparenleft}i{\isacharplus}{\isadigit{1}}{\isacharparenright}{\isacharparenright}\ {\isasymin}\ M{\isacharparenright}{\isacharbraceright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -37,59 +49,76 @@ presentation (see \S\ref{sec:doc-prep-suppress}). In reality one has to define a new datatype and a new function.}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% -{\isachardoublequote}s\ {\isasymTurnstile}\ AF\ f\ \ \ \ {\isacharequal}\ {\isacharparenleft}{\isasymforall}p\ {\isasymin}\ Paths\ s{\isachardot}\ {\isasymexists}i{\isachardot}\ p\ i\ {\isasymTurnstile}\ f{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +{\isachardoublequote}s\ {\isasymTurnstile}\ AF\ f\ \ \ \ {\isacharequal}\ {\isacharparenleft}{\isasymforall}p\ {\isasymin}\ Paths\ s{\isachardot}\ {\isasymexists}i{\isachardot}\ p\ i\ {\isasymTurnstile}\ f{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent Model checking \isa{AF} involves a function which is just complicated enough to warrant a separate definition:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{constdefs}\ af\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}state\ set\ {\isasymRightarrow}\ state\ set\ {\isasymRightarrow}\ state\ set{\isachardoublequote}\isanewline -\ \ \ \ \ \ \ \ \ {\isachardoublequote}af\ A\ T\ {\isasymequiv}\ A\ {\isasymunion}\ {\isacharbraceleft}s{\isachardot}\ {\isasymforall}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\ {\isasymlongrightarrow}\ t\ {\isasymin}\ T{\isacharbraceright}{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ \ \ \ \ {\isachardoublequote}af\ A\ T\ {\isasymequiv}\ A\ {\isasymunion}\ {\isacharbraceleft}s{\isachardot}\ {\isasymforall}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\ {\isasymlongrightarrow}\ t\ {\isasymin}\ T{\isacharbraceright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent Now we define \isa{mc\ {\isacharparenleft}AF\ f{\isacharparenright}} as the least set \isa{T} that includes \isa{mc\ f} and all states all of whose direct successors are in \isa{T}:% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% -{\isachardoublequote}mc{\isacharparenleft}AF\ f{\isacharparenright}\ \ \ \ {\isacharequal}\ lfp{\isacharparenleft}af{\isacharparenleft}mc\ f{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +{\isachardoublequote}mc{\isacharparenleft}AF\ f{\isacharparenright}\ \ \ \ {\isacharequal}\ lfp{\isacharparenleft}af{\isacharparenleft}mc\ f{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent Because \isa{af} is monotone in its second argument (and also its first, but that is irrelevant), \isa{af\ A} has a least fixed point:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ mono{\isacharunderscore}af{\isacharcolon}\ {\isachardoublequote}mono{\isacharparenleft}af\ A{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}simp\ add{\isacharcolon}\ mono{\isacharunderscore}def\ af{\isacharunderscore}def{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}\ blast\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% All we need to prove now is \isa{mc\ {\isacharparenleft}AF\ f{\isacharparenright}\ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ s\ {\isasymTurnstile}\ AF\ f{\isacharbraceright}}, which states @@ -97,9 +126,15 @@ This time we prove the two inclusions separately, starting with the easy one:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{theorem}\ AF{\isacharunderscore}lemma{\isadigit{1}}{\isacharcolon}\isanewline -\ \ {\isachardoublequote}lfp{\isacharparenleft}af\ A{\isacharparenright}\ {\isasymsubseteq}\ {\isacharbraceleft}s{\isachardot}\ {\isasymforall}p\ {\isasymin}\ Paths\ s{\isachardot}\ {\isasymexists}i{\isachardot}\ p\ i\ {\isasymin}\ A{\isacharbraceright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}lfp{\isacharparenleft}af\ A{\isacharparenright}\ {\isasymsubseteq}\ {\isacharbraceleft}s{\isachardot}\ {\isasymforall}p\ {\isasymin}\ Paths\ s{\isachardot}\ {\isasymexists}i{\isachardot}\ p\ i\ {\isasymin}\ A{\isacharbraceright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -112,10 +147,10 @@ The instance of the premise \isa{f\ S\ {\isasymsubseteq}\ S} is proved pointwise, a decision that \isa{auto} takes for us:% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}rule\ lfp{\isacharunderscore}lowerbound{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}auto\ simp\ add{\isacharcolon}\ af{\isacharunderscore}def\ Paths{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}auto\ simp\ add{\isacharcolon}\ af{\isacharunderscore}def\ Paths{\isacharunderscore}def{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -130,12 +165,19 @@ finding the instantiation \isa{{\isasymlambda}i{\isachardot}\ p\ {\isacharparenleft}i\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}} for \isa{{\isasymforall}p}.% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}erule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequote}p\ {\isadigit{1}}{\isachardoublequote}\ \isakeyword{in}\ allE{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% The opposite inclusion is proved by contradiction: if some state @@ -149,9 +191,15 @@ The one-step argument in the sketch above is proved by a variant of contraposition:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ not{\isacharunderscore}in{\isacharunderscore}lfp{\isacharunderscore}afD{\isacharcolon}\isanewline \ {\isachardoublequote}s\ {\isasymnotin}\ lfp{\isacharparenleft}af\ A{\isacharparenright}\ {\isasymLongrightarrow}\ s\ {\isasymnotin}\ A\ {\isasymand}\ {\isacharparenleft}{\isasymexists}\ t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ t\ {\isasymnotin}\ lfp{\isacharparenleft}af\ A{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}erule\ contrapos{\isacharunderscore}np{\isacharparenright}\isanewline \isamarkupfalse% @@ -159,7 +207,14 @@ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}simp\ add{\isacharcolon}\ af{\isacharunderscore}def{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -170,12 +225,12 @@ Now we iterate this process. The following construction of the desired path is parameterized by a predicate \isa{Q} that should hold along the path:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ path\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}state\ {\isasymRightarrow}\ {\isacharparenleft}state\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}nat\ {\isasymRightarrow}\ state{\isacharparenright}{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\isanewline {\isachardoublequote}path\ s\ Q\ {\isadigit{0}}\ {\isacharequal}\ s{\isachardoublequote}\isanewline -{\isachardoublequote}path\ s\ Q\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}SOME\ t{\isachardot}\ {\isacharparenleft}path\ s\ Q\ n{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ Q\ t{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +{\isachardoublequote}path\ s\ Q\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}SOME\ t{\isachardot}\ {\isacharparenleft}path\ s\ Q\ n{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ Q\ t{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -189,35 +244,41 @@ Let us show that if each state \isa{s} that satisfies \isa{Q} has a successor that again satisfies \isa{Q}, then there exists an infinite \isa{Q}-path:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ infinity{\isacharunderscore}lemma{\isacharcolon}\isanewline \ \ {\isachardoublequote}{\isasymlbrakk}\ Q\ s{\isacharsemicolon}\ {\isasymforall}s{\isachardot}\ Q\ s\ {\isasymlongrightarrow}\ {\isacharparenleft}{\isasymexists}\ t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ Q\ t{\isacharparenright}\ {\isasymrbrakk}\ {\isasymLongrightarrow}\isanewline -\ \ \ {\isasymexists}p{\isasymin}Paths\ s{\isachardot}\ {\isasymforall}i{\isachardot}\ Q{\isacharparenleft}p\ i{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ \ {\isasymexists}p{\isasymin}Paths\ s{\isachardot}\ {\isasymforall}i{\isachardot}\ Q{\isacharparenleft}p\ i{\isacharparenright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkuptrue% % \begin{isamarkuptxt}% \noindent First we rephrase the conclusion slightly because we need to prove simultaneously both the path property and the fact that \isa{Q} holds:% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}subgoal{\isacharunderscore}tac\isanewline -\ \ {\isachardoublequote}{\isasymexists}p{\isachardot}\ s\ {\isacharequal}\ p\ {\isadigit{0}}\ {\isasymand}\ {\isacharparenleft}{\isasymforall}i{\isacharcolon}{\isacharcolon}nat{\isachardot}\ {\isacharparenleft}p\ i{\isacharcomma}\ p{\isacharparenleft}i{\isacharplus}{\isadigit{1}}{\isacharparenright}{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ Q{\isacharparenleft}p\ i{\isacharparenright}{\isacharparenright}{\isachardoublequote}{\isacharparenright}\isamarkupfalse% +\ \ {\isachardoublequote}{\isasymexists}p{\isachardot}\ s\ {\isacharequal}\ p\ {\isadigit{0}}\ {\isasymand}\ {\isacharparenleft}{\isasymforall}i{\isacharcolon}{\isacharcolon}nat{\isachardot}\ {\isacharparenleft}p\ i{\isacharcomma}\ p{\isacharparenleft}i{\isacharplus}{\isadigit{1}}{\isacharparenright}{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ Q{\isacharparenleft}p\ i{\isacharparenright}{\isacharparenright}{\isachardoublequote}{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent From this proposition the original goal follows easily:% \end{isamarkuptxt}% -\ \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}simp\ add{\isacharcolon}\ Paths{\isacharunderscore}def{\isacharcomma}\ blast{\isacharparenright}\isamarkupfalse% +\ \isamarkupfalse% +\isacommand{apply}{\isacharparenleft}simp\ add{\isacharcolon}\ Paths{\isacharunderscore}def{\isacharcomma}\ blast{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent The new subgoal is proved by providing the witness \isa{path\ s\ Q} for \isa{p}:% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequote}path\ s\ Q{\isachardoublequote}\ \isakeyword{in}\ exI{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}clarsimp{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}clarsimp{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -229,10 +290,10 @@ \end{isabelle} It invites a proof by induction on \isa{i}:% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ i{\isacharparenright}\isanewline \ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -253,8 +314,8 @@ \isa{{\isacharparenleft}s{\isacharcomma}\ x{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ Q\ x\ {\isasymLongrightarrow}\ {\isacharparenleft}s{\isacharcomma}\ x{\isacharparenright}\ {\isasymin}\ M}, which is trivial. Thus it is not surprising that \isa{fast} can prove the base case quickly:% \end{isamarkuptxt}% -\ \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}fast\ intro{\isacharcolon}\ someI{\isadigit{2}}{\isacharunderscore}ex{\isacharparenright}\isamarkupfalse% +\ \isamarkupfalse% +\isacommand{apply}{\isacharparenleft}fast\ intro{\isacharcolon}\ someI{\isadigit{2}}{\isacharunderscore}ex{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -273,7 +334,7 @@ solve the subgoal and we apply \isa{someI{\isadigit{2}}{\isacharunderscore}ex} by hand. We merely show the proof commands but do not describe the details:% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}{\isacharparenleft}rule\ someI{\isadigit{2}}{\isacharunderscore}ex{\isacharparenright}\isanewline @@ -286,7 +347,14 @@ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}blast{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% Function \isa{path} has fulfilled its purpose now and can be forgotten. @@ -300,40 +368,43 @@ is extensionally equal to \isa{path\ s\ Q}, where \isa{nat{\isacharunderscore}rec} is the predefined primitive recursor on \isa{nat}.% \end{isamarkuptext}% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% % \begin{isamarkuptext}% At last we can prove the opposite direction of \isa{AF{\isacharunderscore}lemma{\isadigit{1}}}:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{theorem}\ AF{\isacharunderscore}lemma{\isadigit{2}}{\isacharcolon}\ {\isachardoublequote}{\isacharbraceleft}s{\isachardot}\ {\isasymforall}p\ {\isasymin}\ Paths\ s{\isachardot}\ {\isasymexists}i{\isachardot}\ p\ i\ {\isasymin}\ A{\isacharbraceright}\ {\isasymsubseteq}\ lfp{\isacharparenleft}af\ A{\isacharparenright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{theorem}\ AF{\isacharunderscore}lemma{\isadigit{2}}{\isacharcolon}\ {\isachardoublequote}{\isacharbraceleft}s{\isachardot}\ {\isasymforall}p\ {\isasymin}\ Paths\ s{\isachardot}\ {\isasymexists}i{\isachardot}\ p\ i\ {\isasymin}\ A{\isacharbraceright}\ {\isasymsubseteq}\ lfp{\isacharparenleft}af\ A{\isacharparenright}{\isachardoublequote}\isamarkupfalse% % \begin{isamarkuptxt}% \noindent The proof is again pointwise and then by contraposition:% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}rule\ subsetI{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}{\isacharparenleft}erule\ contrapos{\isacharunderscore}pp{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{apply}\ simp\isamarkupfalse% +\isacommand{apply}\ simp\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -342,8 +413,8 @@ Applying the \isa{infinity{\isacharunderscore}lemma} as a destruction rule leaves two subgoals, the second premise of \isa{infinity{\isacharunderscore}lemma} and the original subgoal:% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}{\isacharparenleft}drule\ infinity{\isacharunderscore}lemma{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}drule\ infinity{\isacharunderscore}lemma{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -353,10 +424,17 @@ \end{isabelle} Both are solved automatically:% \end{isamarkuptxt}% -\ \isamarkuptrue% +\ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}auto\ dest{\isacharcolon}\ not{\isacharunderscore}in{\isacharunderscore}lfp{\isacharunderscore}afD{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% If you find these proofs too complicated, we recommend that you read @@ -367,14 +445,27 @@ necessary equality \isa{lfp{\isacharparenleft}af\ A{\isacharparenright}\ {\isacharequal}\ {\isachardot}{\isachardot}{\isachardot}} by combining \isa{AF{\isacharunderscore}lemma{\isadigit{1}}} and \isa{AF{\isacharunderscore}lemma{\isadigit{2}}} on the spot:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{theorem}\ {\isachardoublequote}mc\ f\ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ s\ {\isasymTurnstile}\ f{\isacharbraceright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ f{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}{\isacharparenleft}auto\ simp\ add{\isacharcolon}\ EF{\isacharunderscore}lemma\ equalityI{\isacharbrackleft}OF\ AF{\isacharunderscore}lemma{\isadigit{1}}\ AF{\isacharunderscore}lemma{\isadigit{2}}{\isacharbrackright}{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% The language defined above is not quite CTL\@. The latter also includes an @@ -382,13 +473,12 @@ where \isa{f} is true \emph{U}ntil \isa{g} becomes true''. We need an auxiliary function:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ until{\isacharcolon}{\isacharcolon}\ {\isachardoublequote}state\ set\ {\isasymRightarrow}\ state\ set\ {\isasymRightarrow}\ state\ {\isasymRightarrow}\ state\ list\ {\isasymRightarrow}\ bool{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\isanewline {\isachardoublequote}until\ A\ B\ s\ {\isacharbrackleft}{\isacharbrackright}\ \ \ \ {\isacharequal}\ {\isacharparenleft}s\ {\isasymin}\ B{\isacharparenright}{\isachardoublequote}\isanewline -{\isachardoublequote}until\ A\ B\ s\ {\isacharparenleft}t{\isacharhash}p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}s\ {\isasymin}\ A\ {\isasymand}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ until\ A\ B\ t\ p{\isacharparenright}{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% +{\isachardoublequote}until\ A\ B\ s\ {\isacharparenleft}t{\isacharhash}p{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}s\ {\isasymin}\ A\ {\isasymand}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ until\ A\ B\ t\ p{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -415,32 +505,46 @@ \end{exercise} For more CTL exercises see, for example, Huth and Ryan \cite{Huth-Ryan-book}.% \end{isamarkuptext}% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% % \begin{isamarkuptext}% Let us close this section with a few words about the executability of @@ -455,8 +559,19 @@ from HOL definitions, but that is beyond the scope of the tutorial.% \index{CTL|)}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/CTL/document/CTLind.tex --- a/doc-src/TutorialI/CTL/document/CTLind.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/CTL/document/CTLind.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{CTLind}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsubsection{CTL Revisited% } @@ -28,12 +41,12 @@ % Second proof of opposite direction, directly by well-founded induction % on the initial segment of M that avoids A.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ Avoid\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}state\ {\isasymRightarrow}\ state\ set\ {\isasymRightarrow}\ state\ set{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{inductive}\ {\isachardoublequote}Avoid\ s\ A{\isachardoublequote}\isanewline \isakeyword{intros}\ {\isachardoublequote}s\ {\isasymin}\ Avoid\ s\ A{\isachardoublequote}\isanewline -\ \ \ \ \ \ \ {\isachardoublequote}{\isasymlbrakk}\ t\ {\isasymin}\ Avoid\ s\ A{\isacharsemicolon}\ t\ {\isasymnotin}\ A{\isacharsemicolon}\ {\isacharparenleft}t{\isacharcomma}u{\isacharparenright}\ {\isasymin}\ M\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ u\ {\isasymin}\ Avoid\ s\ A{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ \ \ {\isachardoublequote}{\isasymlbrakk}\ t\ {\isasymin}\ Avoid\ s\ A{\isacharsemicolon}\ t\ {\isasymnotin}\ A{\isacharsemicolon}\ {\isacharparenleft}t{\isacharcomma}u{\isacharparenright}\ {\isasymin}\ M\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ u\ {\isasymin}\ Avoid\ s\ A{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% It is easy to see that for any infinite \isa{A}-avoiding path \isa{f} @@ -45,10 +58,16 @@ reformulation, as explained in \S\ref{sec:ind-var-in-prems} above; the \isa{rule{\isacharunderscore}format} directive undoes the reformulation after the proof.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ ex{\isacharunderscore}infinite{\isacharunderscore}path{\isacharbrackleft}rule{\isacharunderscore}format{\isacharbrackright}{\isacharcolon}\isanewline \ \ {\isachardoublequote}t\ {\isasymin}\ Avoid\ s\ A\ \ {\isasymLongrightarrow}\isanewline \ \ \ {\isasymforall}f{\isasymin}Paths\ t{\isachardot}\ {\isacharparenleft}{\isasymforall}i{\isachardot}\ f\ i\ {\isasymnotin}\ A{\isacharparenright}\ {\isasymlongrightarrow}\ {\isacharparenleft}{\isasymexists}p{\isasymin}Paths\ s{\isachardot}\ {\isasymforall}i{\isachardot}\ p\ i\ {\isasymnotin}\ A{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}erule\ Avoid{\isachardot}induct{\isacharparenright}\isanewline \ \isamarkupfalse% @@ -60,7 +79,14 @@ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ Paths{\isacharunderscore}def\ split{\isacharcolon}\ nat{\isachardot}split{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -78,9 +104,15 @@ ``between'' \isa{s} and \isa{A}, in other words all of \isa{Avoid\ s\ A}, is contained in \isa{lfp\ {\isacharparenleft}af\ A{\isacharparenright}}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ Avoid{\isacharunderscore}in{\isacharunderscore}lfp{\isacharbrackleft}rule{\isacharunderscore}format{\isacharparenleft}no{\isacharunderscore}asm{\isacharparenright}{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequote}{\isasymforall}p{\isasymin}Paths\ s{\isachardot}\ {\isasymexists}i{\isachardot}\ p\ i\ {\isasymin}\ A\ {\isasymLongrightarrow}\ t\ {\isasymin}\ Avoid\ s\ A\ {\isasymlongrightarrow}\ t\ {\isasymin}\ lfp{\isacharparenleft}af\ A{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}{\isasymforall}p{\isasymin}Paths\ s{\isachardot}\ {\isasymexists}i{\isachardot}\ p\ i\ {\isasymin}\ A\ {\isasymLongrightarrow}\ t\ {\isasymin}\ Avoid\ s\ A\ {\isasymlongrightarrow}\ t\ {\isasymin}\ lfp{\isacharparenleft}af\ A{\isacharparenright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -99,13 +131,12 @@ starting from \isa{s} implies well-foundedness of this relation. For the moment we assume this and proceed with the induction:% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}subgoal{\isacharunderscore}tac\ {\isachardoublequote}wf{\isacharbraceleft}{\isacharparenleft}y{\isacharcomma}x{\isacharparenright}{\isachardot}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ M\ {\isasymand}\ x\ {\isasymin}\ Avoid\ s\ A\ {\isasymand}\ x\ {\isasymnotin}\ A{\isacharbraceright}{\isachardoublequote}{\isacharparenright}\isanewline \ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}erule{\isacharunderscore}tac\ a\ {\isacharequal}\ t\ \isakeyword{in}\ wf{\isacharunderscore}induct{\isacharparenright}\isanewline \ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}clarsimp{\isacharparenright}\isamarkupfalse% -\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}clarsimp{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -129,12 +160,12 @@ Hence, by the induction hypothesis, all successors of \isa{t} are indeed in \isa{lfp\ {\isacharparenleft}af\ A{\isacharparenright}}. Mechanically:% \end{isamarkuptxt}% -\ \isamarkuptrue% +\ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}subst\ lfp{\isacharunderscore}unfold{\isacharbrackleft}OF\ mono{\isacharunderscore}af{\isacharbrackright}{\isacharparenright}\isanewline \ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}simp\ {\isacharparenleft}no{\isacharunderscore}asm{\isacharparenright}\ add{\isacharcolon}\ af{\isacharunderscore}def{\isacharparenright}\isanewline \ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ Avoid{\isachardot}intros{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ Avoid{\isachardot}intros{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% Having proved the main goal, we return to the proof obligation that the @@ -147,7 +178,7 @@ From lemma \isa{ex{\isacharunderscore}infinite{\isacharunderscore}path} the existence of an infinite \isa{A}-avoiding path starting in \isa{s} follows, contradiction.% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}erule\ contrapos{\isacharunderscore}pp{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}{\isacharparenleft}simp\ add{\isacharcolon}\ wf{\isacharunderscore}iff{\isacharunderscore}no{\isacharunderscore}infinite{\isacharunderscore}down{\isacharunderscore}chain{\isacharparenright}\isanewline @@ -158,7 +189,14 @@ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}auto\ simp\ add{\isacharcolon}\ Paths{\isacharunderscore}def{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% The \isa{{\isacharparenleft}no{\isacharunderscore}asm{\isacharparenright}} modifier of the \isa{rule{\isacharunderscore}format} directive in the @@ -175,13 +213,37 @@ by the first \isa{Avoid}-rule. Isabelle confirms this:% \index{CTL|)}% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{theorem}\ AF{\isacharunderscore}lemma{\isadigit{2}}{\isacharcolon}\ \ {\isachardoublequote}{\isacharbraceleft}s{\isachardot}\ {\isasymforall}p\ {\isasymin}\ Paths\ s{\isachardot}\ {\isasymexists}\ i{\isachardot}\ p\ i\ {\isasymin}\ A{\isacharbraceright}\ {\isasymsubseteq}\ lfp{\isacharparenleft}af\ A{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{by}{\isacharparenleft}auto\ elim{\isacharcolon}\ Avoid{\isacharunderscore}in{\isacharunderscore}lfp\ intro{\isacharcolon}\ Avoid{\isachardot}intros{\isacharparenright}\isanewline \isanewline -\isamarkupfalse% -\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/CTL/document/PDL.tex --- a/doc-src/TutorialI/CTL/document/PDL.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/CTL/document/PDL.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{PDL}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsubsection{Propositional Dynamic Logic --- PDL% } @@ -17,12 +30,12 @@ \cite{HarelKT-DL} looks quite different from ours, but the two are easily shown to be equivalent.}% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{datatype}\ formula\ {\isacharequal}\ Atom\ atom\isanewline \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ Neg\ formula\isanewline \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ And\ formula\ formula\isanewline \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ AX\ formula\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ EF\ formula\isamarkupfalse% +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ EF\ formula\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -31,8 +44,8 @@ A validity relation between states and formulae specifies the semantics:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{consts}\ valid\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}state\ {\isasymRightarrow}\ formula\ {\isasymRightarrow}\ bool{\isachardoublequote}\ \ \ {\isacharparenleft}{\isachardoublequote}{\isacharparenleft}{\isacharunderscore}\ {\isasymTurnstile}\ {\isacharunderscore}{\isacharparenright}{\isachardoublequote}\ {\isacharbrackleft}{\isadigit{8}}{\isadigit{0}}{\isacharcomma}{\isadigit{8}}{\isadigit{0}}{\isacharbrackright}\ {\isadigit{8}}{\isadigit{0}}{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{consts}\ valid\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}state\ {\isasymRightarrow}\ formula\ {\isasymRightarrow}\ bool{\isachardoublequote}\ \ \ {\isacharparenleft}{\isachardoublequote}{\isacharparenleft}{\isacharunderscore}\ {\isasymTurnstile}\ {\isacharunderscore}{\isacharparenright}{\isachardoublequote}\ {\isacharbrackleft}{\isadigit{8}}{\isadigit{0}}{\isacharcomma}{\isadigit{8}}{\isadigit{0}}{\isacharbrackright}\ {\isadigit{8}}{\isadigit{0}}{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -40,13 +53,13 @@ \hbox{\isa{valid\ s\ f}}. The definition of \isa{{\isasymTurnstile}} is by recursion over the syntax:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{primrec}\isanewline {\isachardoublequote}s\ {\isasymTurnstile}\ Atom\ a\ \ {\isacharequal}\ {\isacharparenleft}a\ {\isasymin}\ L\ s{\isacharparenright}{\isachardoublequote}\isanewline {\isachardoublequote}s\ {\isasymTurnstile}\ Neg\ f\ \ \ {\isacharequal}\ {\isacharparenleft}{\isasymnot}{\isacharparenleft}s\ {\isasymTurnstile}\ f{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isanewline {\isachardoublequote}s\ {\isasymTurnstile}\ And\ f\ g\ {\isacharequal}\ {\isacharparenleft}s\ {\isasymTurnstile}\ f\ {\isasymand}\ s\ {\isasymTurnstile}\ g{\isacharparenright}{\isachardoublequote}\isanewline {\isachardoublequote}s\ {\isasymTurnstile}\ AX\ f\ \ \ \ {\isacharequal}\ {\isacharparenleft}{\isasymforall}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\ {\isasymlongrightarrow}\ t\ {\isasymTurnstile}\ f{\isacharparenright}{\isachardoublequote}\isanewline -{\isachardoublequote}s\ {\isasymTurnstile}\ EF\ f\ \ \ \ {\isacharequal}\ {\isacharparenleft}{\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}\ {\isasymand}\ t\ {\isasymTurnstile}\ f{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +{\isachardoublequote}s\ {\isasymTurnstile}\ EF\ f\ \ \ \ {\isacharequal}\ {\isacharparenleft}{\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}\ {\isasymand}\ t\ {\isasymTurnstile}\ f{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -59,7 +72,7 @@ Now we come to the model checker itself. It maps a formula into the set of states where the formula is true. It too is defined by recursion over the syntax:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ mc\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}formula\ {\isasymRightarrow}\ state\ set{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\isanewline @@ -67,7 +80,7 @@ {\isachardoublequote}mc{\isacharparenleft}Neg\ f{\isacharparenright}\ \ \ {\isacharequal}\ {\isacharminus}mc\ f{\isachardoublequote}\isanewline {\isachardoublequote}mc{\isacharparenleft}And\ f\ g{\isacharparenright}\ {\isacharequal}\ mc\ f\ {\isasyminter}\ mc\ g{\isachardoublequote}\isanewline {\isachardoublequote}mc{\isacharparenleft}AX\ f{\isacharparenright}\ \ \ \ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ {\isasymforall}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\ \ {\isasymlongrightarrow}\ t\ {\isasymin}\ mc\ f{\isacharbraceright}{\isachardoublequote}\isanewline -{\isachardoublequote}mc{\isacharparenleft}EF\ f{\isacharparenright}\ \ \ \ {\isacharequal}\ lfp{\isacharparenleft}{\isasymlambda}T{\isachardot}\ mc\ f\ {\isasymunion}\ {\isacharparenleft}M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +{\isachardoublequote}mc{\isacharparenleft}EF\ f{\isacharparenright}\ \ \ \ {\isacharequal}\ lfp{\isacharparenleft}{\isasymlambda}T{\isachardot}\ mc\ f\ {\isasymunion}\ {\isacharparenleft}M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -84,36 +97,54 @@ First we prove monotonicity of the function inside \isa{lfp} in order to make sure it really has a least fixed point.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ mono{\isacharunderscore}ef{\isacharcolon}\ {\isachardoublequote}mono{\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ {\isacharparenleft}M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}rule\ monoI{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}\ blast\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Now we can relate model checking and semantics. For the \isa{EF} case we need a separate lemma:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ EF{\isacharunderscore}lemma{\isacharcolon}\isanewline -\ \ {\isachardoublequote}lfp{\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ {\isacharparenleft}M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A{\isacharbraceright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}lfp{\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ {\isacharparenleft}M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}\ T{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A{\isacharbraceright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkuptrue% % \begin{isamarkuptxt}% \noindent The equality is proved in the canonical fashion by proving that each set includes the other; the inclusion is shown pointwise:% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}rule\ equalityI{\isacharparenright}\isanewline \ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}rule\ subsetI{\isacharparenright}\isanewline \ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isamarkupfalse% -\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -123,12 +154,12 @@ \end{isabelle} which is proved by \isa{lfp}-induction:% \end{isamarkuptxt}% -\ \isamarkuptrue% +\ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}erule\ lfp{\isacharunderscore}induct{\isacharparenright}\isanewline \ \ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}rule\ mono{\isacharunderscore}ef{\isacharparenright}\isanewline \ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -142,17 +173,17 @@ It is proved by \isa{blast}, using the transitivity of \isa{M\isactrlsup {\isacharasterisk}}.% \end{isamarkuptxt}% -\ \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ rtrancl{\isacharunderscore}trans{\isacharparenright}\isamarkupfalse% +\ \isamarkupfalse% +\isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ rtrancl{\isacharunderscore}trans{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% We now return to the second set inclusion subgoal, which is again proved pointwise:% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}rule\ subsetI{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp{\isacharcomma}\ clarify{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}simp{\isacharcomma}\ clarify{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -174,8 +205,8 @@ \isa{P\ a} provided each step backwards from a predecessor \isa{z} of \isa{b} preserves \isa{P}.% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}{\isacharparenleft}erule\ converse{\isacharunderscore}rtrancl{\isacharunderscore}induct{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}erule\ converse{\isacharunderscore}rtrancl{\isacharunderscore}induct{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -185,8 +216,8 @@ \end{isabelle} is solved by unrolling \isa{lfp} once% \end{isamarkuptxt}% -\ \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}subst\ lfp{\isacharunderscore}unfold{\isacharbrackleft}OF\ mono{\isacharunderscore}ef{\isacharbrackright}{\isacharparenright}\isamarkupfalse% +\ \isamarkupfalse% +\isacommand{apply}{\isacharparenleft}subst\ lfp{\isacharunderscore}unfold{\isacharbrackleft}OF\ mono{\isacharunderscore}ef{\isacharbrackright}{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -194,32 +225,52 @@ \end{isabelle} and disposing of the resulting trivial subgoal automatically:% \end{isamarkuptxt}% -\ \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}blast{\isacharparenright}\isamarkupfalse% +\ \isamarkupfalse% +\isacommand{apply}{\isacharparenleft}blast{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent The proof of the induction step is identical to the one for the base case:% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}subst\ lfp{\isacharunderscore}unfold{\isacharbrackleft}OF\ mono{\isacharunderscore}ef{\isacharbrackright}{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}{\isacharparenleft}blast{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% The main theorem is proved in the familiar manner: induction followed by \isa{auto} augmented with the lemma as a simplification rule.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{theorem}\ {\isachardoublequote}mc\ f\ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ s\ {\isasymTurnstile}\ f{\isacharbraceright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ f{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}{\isacharparenleft}auto\ simp\ add{\isacharcolon}\ EF{\isacharunderscore}lemma{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \begin{exercise} @@ -240,20 +291,58 @@ \end{exercise} \index{PDL|)}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/CodeGen/document/CodeGen.tex --- a/doc-src/TutorialI/CodeGen/document/CodeGen.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/CodeGen/document/CodeGen.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{CodeGen}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsection{Case Study: Compiling Expressions% } @@ -18,12 +31,12 @@ a fixed set of binary operations: instead the expression contains the appropriate function itself.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{types}\ {\isacharprime}v\ binop\ {\isacharequal}\ {\isachardoublequote}{\isacharprime}v\ {\isasymRightarrow}\ {\isacharprime}v\ {\isasymRightarrow}\ {\isacharprime}v{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{datatype}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}expr\ {\isacharequal}\ Cex\ {\isacharprime}v\isanewline \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ Vex\ {\isacharprime}a\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ Bex\ {\isachardoublequote}{\isacharprime}v\ binop{\isachardoublequote}\ \ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}expr{\isachardoublequote}\ \ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}expr{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ Bex\ {\isachardoublequote}{\isacharprime}v\ binop{\isachardoublequote}\ \ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}expr{\isachardoublequote}\ \ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}expr{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -33,13 +46,13 @@ The value of an expression with respect to an environment that maps variables to values is easily defined:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ value\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}expr\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}v{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}v{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\isanewline {\isachardoublequote}value\ {\isacharparenleft}Cex\ v{\isacharparenright}\ env\ {\isacharequal}\ v{\isachardoublequote}\isanewline {\isachardoublequote}value\ {\isacharparenleft}Vex\ a{\isacharparenright}\ env\ {\isacharequal}\ env\ a{\isachardoublequote}\isanewline -{\isachardoublequote}value\ {\isacharparenleft}Bex\ f\ e{\isadigit{1}}\ e{\isadigit{2}}{\isacharparenright}\ env\ {\isacharequal}\ f\ {\isacharparenleft}value\ e{\isadigit{1}}\ env{\isacharparenright}\ {\isacharparenleft}value\ e{\isadigit{2}}\ env{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +{\isachardoublequote}value\ {\isacharparenleft}Bex\ f\ e{\isadigit{1}}\ e{\isadigit{2}}{\isacharparenright}\ env\ {\isacharequal}\ f\ {\isacharparenleft}value\ e{\isadigit{1}}\ env{\isacharparenright}\ {\isacharparenleft}value\ e{\isadigit{2}}\ env{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% The stack machine has three instructions: load a constant value onto the @@ -47,10 +60,10 @@ binary operation to the two topmost elements of the stack, replacing them by the result. As for \isa{expr}, addresses and values are type parameters:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{datatype}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}\ instr\ {\isacharequal}\ Const\ {\isacharprime}v\isanewline \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ Load\ {\isacharprime}a\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ Apply\ {\isachardoublequote}{\isacharprime}v\ binop{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ Apply\ {\isachardoublequote}{\isacharprime}v\ binop{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% The execution of the stack machine is modelled by a function @@ -60,7 +73,7 @@ and returns the stack at the end of the execution --- the store remains unchanged:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ exec\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}instr\ list\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a{\isasymRightarrow}{\isacharprime}v{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}v\ list\ {\isasymRightarrow}\ {\isacharprime}v\ list{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\isanewline @@ -68,7 +81,7 @@ {\isachardoublequote}exec\ {\isacharparenleft}i{\isacharhash}is{\isacharparenright}\ s\ vs\ {\isacharequal}\ {\isacharparenleft}case\ i\ of\isanewline \ \ \ \ Const\ v\ \ {\isasymRightarrow}\ exec\ is\ s\ {\isacharparenleft}v{\isacharhash}vs{\isacharparenright}\isanewline \ \ {\isacharbar}\ Load\ a\ \ \ {\isasymRightarrow}\ exec\ is\ s\ {\isacharparenleft}{\isacharparenleft}s\ a{\isacharparenright}{\isacharhash}vs{\isacharparenright}\isanewline -\ \ {\isacharbar}\ Apply\ f\ \ {\isasymRightarrow}\ exec\ is\ s\ {\isacharparenleft}{\isacharparenleft}f\ {\isacharparenleft}hd\ vs{\isacharparenright}\ {\isacharparenleft}hd{\isacharparenleft}tl\ vs{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharhash}{\isacharparenleft}tl{\isacharparenleft}tl\ vs{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isacharbar}\ Apply\ f\ \ {\isasymRightarrow}\ exec\ is\ s\ {\isacharparenleft}{\isacharparenleft}f\ {\isacharparenleft}hd\ vs{\isacharparenright}\ {\isacharparenleft}hd{\isacharparenleft}tl\ vs{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharhash}{\isacharparenleft}tl{\isacharparenleft}tl\ vs{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -83,28 +96,46 @@ The compiler is a function from expressions to a list of instructions. Its definition is obvious:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ comp\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}expr\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}instr\ list{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\isanewline {\isachardoublequote}comp\ {\isacharparenleft}Cex\ v{\isacharparenright}\ \ \ \ \ \ \ {\isacharequal}\ {\isacharbrackleft}Const\ v{\isacharbrackright}{\isachardoublequote}\isanewline {\isachardoublequote}comp\ {\isacharparenleft}Vex\ a{\isacharparenright}\ \ \ \ \ \ \ {\isacharequal}\ {\isacharbrackleft}Load\ a{\isacharbrackright}{\isachardoublequote}\isanewline -{\isachardoublequote}comp\ {\isacharparenleft}Bex\ f\ e{\isadigit{1}}\ e{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}comp\ e{\isadigit{2}}{\isacharparenright}\ {\isacharat}\ {\isacharparenleft}comp\ e{\isadigit{1}}{\isacharparenright}\ {\isacharat}\ {\isacharbrackleft}Apply\ f{\isacharbrackright}{\isachardoublequote}\isamarkupfalse% +{\isachardoublequote}comp\ {\isacharparenleft}Bex\ f\ e{\isadigit{1}}\ e{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}comp\ e{\isadigit{2}}{\isacharparenright}\ {\isacharat}\ {\isacharparenleft}comp\ e{\isadigit{1}}{\isacharparenright}\ {\isacharat}\ {\isacharbrackleft}Apply\ f{\isacharbrackright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% Now we have to prove the correctness of the compiler, i.e.\ that the execution of a compiled expression results in the value of the expression:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{theorem}\ {\isachardoublequote}exec\ {\isacharparenleft}comp\ e{\isacharparenright}\ s\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}value\ e\ s{\isacharbrackright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{theorem}\ {\isachardoublequote}exec\ {\isacharparenleft}comp\ e{\isacharparenright}\ s\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}value\ e\ s{\isacharbrackright}{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% % \begin{isamarkuptext}% \noindent This theorem needs to be generalized:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{theorem}\ {\isachardoublequote}{\isasymforall}vs{\isachardot}\ exec\ {\isacharparenleft}comp\ e{\isacharparenright}\ s\ vs\ {\isacharequal}\ {\isacharparenleft}value\ e\ s{\isacharparenright}\ {\isacharhash}\ vs{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{theorem}\ {\isachardoublequote}{\isasymforall}vs{\isachardot}\ exec\ {\isacharparenleft}comp\ e{\isacharparenright}\ s\ vs\ {\isacharequal}\ {\isacharparenleft}value\ e\ s{\isacharparenright}\ {\isacharhash}\ vs{\isachardoublequote}\isamarkupfalse% % \begin{isamarkuptxt}% \noindent @@ -112,10 +143,22 @@ First, we must prove a lemma about executing the concatenation of two instruction sequences:% \end{isamarkuptxt}% -\isamarkuptrue% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkupfalse% \isacommand{lemma}\ exec{\isacharunderscore}app{\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequote}{\isasymforall}vs{\isachardot}\ exec\ {\isacharparenleft}xs{\isacharat}ys{\isacharparenright}\ s\ vs\ {\isacharequal}\ exec\ ys\ s\ {\isacharparenleft}exec\ xs\ s\ vs{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}{\isasymforall}vs{\isachardot}\ exec\ {\isacharparenleft}xs{\isacharat}ys{\isacharparenright}\ s\ vs\ {\isacharequal}\ exec\ ys\ s\ {\isacharparenleft}exec\ xs\ s\ vs{\isacharparenright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -124,9 +167,15 @@ that contains two \isa{case}-expressions over instructions. Thus we add automatic case splitting, which finishes the proof:% \end{isamarkuptxt}% +\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharcomma}\ simp{\isacharcomma}\ simp\ split{\isacharcolon}\ instr{\isachardot}split{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharcomma}\ simp{\isacharcomma}\ simp\ split{\isacharcolon}\ instr{\isachardot}split{\isacharparenright}\isamarkupfalse% -\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -134,11 +183,21 @@ be modified in the same way as \isa{simp}. Thus the proof can be rewritten as% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharcomma}\ simp{\isacharunderscore}all\ split{\isacharcolon}\ instr{\isachardot}split{\isacharparenright}\isamarkupfalse% -\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharcomma}\ simp{\isacharunderscore}all\ split{\isacharcolon}\ instr{\isachardot}split{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -150,10 +209,32 @@ its instance.% \index{compiling expressions example|)}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Datatype/document/ABexpr.tex --- a/doc-src/TutorialI/Datatype/document/ABexpr.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Datatype/document/ABexpr.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{ABexpr}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% \index{datatypes!mutually recursive}% @@ -17,7 +30,7 @@ \end{itemize} In Isabelle this becomes% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{datatype}\ {\isacharprime}a\ aexp\ {\isacharequal}\ IF\ \ \ {\isachardoublequote}{\isacharprime}a\ bexp{\isachardoublequote}\ {\isachardoublequote}{\isacharprime}a\ aexp{\isachardoublequote}\ {\isachardoublequote}{\isacharprime}a\ aexp{\isachardoublequote}\isanewline \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ Sum\ \ {\isachardoublequote}{\isacharprime}a\ aexp{\isachardoublequote}\ {\isachardoublequote}{\isacharprime}a\ aexp{\isachardoublequote}\isanewline \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ Diff\ {\isachardoublequote}{\isacharprime}a\ aexp{\isachardoublequote}\ {\isachardoublequote}{\isacharprime}a\ aexp{\isachardoublequote}\isanewline @@ -25,7 +38,7 @@ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ Num\ nat\isanewline \isakeyword{and}\ \ \ \ \ \ {\isacharprime}a\ bexp\ {\isacharequal}\ Less\ {\isachardoublequote}{\isacharprime}a\ aexp{\isachardoublequote}\ {\isachardoublequote}{\isacharprime}a\ aexp{\isachardoublequote}\isanewline \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ And\ \ {\isachardoublequote}{\isacharprime}a\ bexp{\isachardoublequote}\ {\isachardoublequote}{\isacharprime}a\ bexp{\isachardoublequote}\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ Neg\ \ {\isachardoublequote}{\isacharprime}a\ bexp{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ Neg\ \ {\isachardoublequote}{\isacharprime}a\ bexp{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -36,9 +49,9 @@ expressions can be arithmetic comparisons, conjunctions and negations. The semantics is given by two evaluation functions:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ \ evala\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ aexp\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ nat{\isachardoublequote}\isanewline -\ \ \ \ \ \ \ \ evalb\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ bexp\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ bool{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ \ \ \ evalb\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ bexp\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ nat{\isacharparenright}\ {\isasymRightarrow}\ bool{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -48,7 +61,7 @@ operate on them. Hence they need to be defined in a single \isacommand{primrec} section:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{primrec}\isanewline \ \ {\isachardoublequote}evala\ {\isacharparenleft}IF\ b\ a{\isadigit{1}}\ a{\isadigit{2}}{\isacharparenright}\ env\ {\isacharequal}\isanewline \ \ \ \ \ {\isacharparenleft}if\ evalb\ b\ env\ then\ evala\ a{\isadigit{1}}\ env\ else\ evala\ a{\isadigit{2}}\ env{\isacharparenright}{\isachardoublequote}\isanewline @@ -59,15 +72,15 @@ \isanewline \ \ {\isachardoublequote}evalb\ {\isacharparenleft}Less\ a{\isadigit{1}}\ a{\isadigit{2}}{\isacharparenright}\ env\ {\isacharequal}\ {\isacharparenleft}evala\ a{\isadigit{1}}\ env\ {\isacharless}\ evala\ a{\isadigit{2}}\ env{\isacharparenright}{\isachardoublequote}\isanewline \ \ {\isachardoublequote}evalb\ {\isacharparenleft}And\ b{\isadigit{1}}\ b{\isadigit{2}}{\isacharparenright}\ env\ {\isacharequal}\ {\isacharparenleft}evalb\ b{\isadigit{1}}\ env\ {\isasymand}\ evalb\ b{\isadigit{2}}\ env{\isacharparenright}{\isachardoublequote}\isanewline -\ \ {\isachardoublequote}evalb\ {\isacharparenleft}Neg\ b{\isacharparenright}\ env\ {\isacharequal}\ {\isacharparenleft}{\isasymnot}\ evalb\ b\ env{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}evalb\ {\isacharparenleft}Neg\ b{\isacharparenright}\ env\ {\isacharequal}\ {\isacharparenleft}{\isasymnot}\ evalb\ b\ env{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent In the same fashion we also define two functions that perform substitution:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ substa\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ aexp{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ aexp\ {\isasymRightarrow}\ {\isacharprime}b\ aexp{\isachardoublequote}\isanewline -\ \ \ \ \ \ \ substb\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ aexp{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ bexp\ {\isasymRightarrow}\ {\isacharprime}b\ bexp{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ \ \ substb\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ aexp{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ bexp\ {\isasymRightarrow}\ {\isacharprime}b\ bexp{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -76,7 +89,7 @@ result, the type of variables in the expression may change from \isa{{\isacharprime}a} to \isa{{\isacharprime}b}. Note that there are only arithmetic and no boolean variables.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{primrec}\isanewline \ \ {\isachardoublequote}substa\ s\ {\isacharparenleft}IF\ b\ a{\isadigit{1}}\ a{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\isanewline \ \ \ \ \ IF\ {\isacharparenleft}substb\ s\ b{\isacharparenright}\ {\isacharparenleft}substa\ s\ a{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}substa\ s\ a{\isadigit{2}}{\isacharparenright}{\isachardoublequote}\isanewline @@ -87,7 +100,7 @@ \isanewline \ \ {\isachardoublequote}substb\ s\ {\isacharparenleft}Less\ a{\isadigit{1}}\ a{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ Less\ {\isacharparenleft}substa\ s\ a{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}substa\ s\ a{\isadigit{2}}{\isacharparenright}{\isachardoublequote}\isanewline \ \ {\isachardoublequote}substb\ s\ {\isacharparenleft}And\ b{\isadigit{1}}\ b{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ And\ {\isacharparenleft}substb\ s\ b{\isadigit{1}}{\isacharparenright}\ {\isacharparenleft}substb\ s\ b{\isadigit{2}}{\isacharparenright}{\isachardoublequote}\isanewline -\ \ {\isachardoublequote}substb\ s\ {\isacharparenleft}Neg\ b{\isacharparenright}\ {\isacharequal}\ Neg\ {\isacharparenleft}substb\ s\ b{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}substb\ s\ {\isacharparenleft}Neg\ b{\isacharparenright}\ {\isacharequal}\ Neg\ {\isacharparenleft}substb\ s\ b{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% Now we can prove a fundamental theorem about the interaction between @@ -99,19 +112,31 @@ theorem in the induction step. Therefore you need to state and prove both theorems simultaneously:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}evala\ {\isacharparenleft}substa\ s\ a{\isacharparenright}\ env\ {\isacharequal}\ evala\ a\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ evala\ {\isacharparenleft}s\ x{\isacharparenright}\ env{\isacharparenright}\ {\isasymand}\isanewline \ \ \ \ \ \ \ \ evalb\ {\isacharparenleft}substb\ s\ b{\isacharparenright}\ env\ {\isacharequal}\ evalb\ b\ {\isacharparenleft}{\isasymlambda}x{\isachardot}\ evala\ {\isacharparenleft}s\ x{\isacharparenright}\ env{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ a\ \isakeyword{and}\ b{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ a\ \isakeyword{and}\ b{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent The resulting 8 goals (one for each constructor) are proved in one fell swoop:% \end{isamarkuptxt}% +\isamarkupfalse% +\isacommand{apply}\ simp{\isacharunderscore}all% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{apply}\ simp{\isacharunderscore}all\isamarkupfalse% -\isamarkupfalse% % \begin{isamarkuptext}% In general, given $n$ mutually recursive datatypes $\tau@1$, \dots, $\tau@n$, @@ -134,20 +159,45 @@ of type annotations following lemma \isa{subst{\isacharunderscore}id} below). \end{exercise}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Datatype/document/Fundata.tex --- a/doc-src/TutorialI/Datatype/document/Fundata.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Datatype/document/Fundata.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,8 +1,21 @@ % \begin{isabellebody}% \def\isabellecontext{Fundata}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \isamarkupfalse% -\isacommand{datatype}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}i{\isacharparenright}bigtree\ {\isacharequal}\ Tip\ {\isacharbar}\ Br\ {\isacharprime}a\ {\isachardoublequote}{\isacharprime}i\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}i{\isacharparenright}bigtree{\isachardoublequote}\isamarkupfalse% +\isacommand{datatype}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}i{\isacharparenright}bigtree\ {\isacharequal}\ Tip\ {\isacharbar}\ Br\ {\isacharprime}a\ {\isachardoublequote}{\isacharprime}i\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}i{\isacharparenright}bigtree{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -22,12 +35,12 @@ Function \isa{map{\isacharunderscore}bt} applies a function to all labels in a \isa{bigtree}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ map{\isacharunderscore}bt\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}i{\isacharparenright}bigtree\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}b{\isacharcomma}{\isacharprime}i{\isacharparenright}bigtree{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\isanewline {\isachardoublequote}map{\isacharunderscore}bt\ f\ Tip\ \ \ \ \ \ {\isacharequal}\ Tip{\isachardoublequote}\isanewline -{\isachardoublequote}map{\isacharunderscore}bt\ f\ {\isacharparenleft}Br\ a\ F{\isacharparenright}\ {\isacharequal}\ Br\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}{\isasymlambda}i{\isachardot}\ map{\isacharunderscore}bt\ f\ {\isacharparenleft}F\ i{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +{\isachardoublequote}map{\isacharunderscore}bt\ f\ {\isacharparenleft}Br\ a\ F{\isacharparenright}\ {\isacharequal}\ Br\ {\isacharparenleft}f\ a{\isacharparenright}\ {\isacharparenleft}{\isasymlambda}i{\isachardot}\ map{\isacharunderscore}bt\ f\ {\isacharparenleft}F\ i{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent This is a valid \isacommand{primrec} definition because the @@ -40,14 +53,31 @@ The following lemma has a simple proof by induction:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}map{\isacharunderscore}bt\ {\isacharparenleft}g\ o\ f{\isacharparenright}\ T\ {\isacharequal}\ map{\isacharunderscore}bt\ g\ {\isacharparenleft}map{\isacharunderscore}bt\ f\ T{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ T{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -59,9 +89,26 @@ \isaindent{\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}a\ F{\isachardot}\ }map{\isacharunderscore}bt\ {\isacharparenleft}g\ {\isasymcirc}\ f{\isacharparenright}\ {\isacharparenleft}Br\ a\ F{\isacharparenright}\ {\isacharequal}\ map{\isacharunderscore}bt\ g\ {\isacharparenleft}map{\isacharunderscore}bt\ f\ {\isacharparenleft}Br\ a\ F{\isacharparenright}{\isacharparenright}% \end{isabelle}% \end{isamarkuptxt}% -\isamarkuptrue% -\isamarkupfalse% -\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Datatype/document/Nested.tex --- a/doc-src/TutorialI/Datatype/document/Nested.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Datatype/document/Nested.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Nested}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% \index{datatypes!and nested recursion}% @@ -12,9 +25,8 @@ Consider the following model of terms where function symbols can be applied to a list of arguments:% \end{isamarkuptext}% -\isamarkuptrue% \isamarkupfalse% -\isacommand{datatype}\ {\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}{\isachardoublequote}term{\isachardoublequote}\ {\isacharequal}\ Var\ {\isacharprime}v\ {\isacharbar}\ App\ {\isacharprime}f\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}term\ list{\isachardoublequote}\isamarkupfalse% +\isacommand{datatype}\ {\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}{\isachardoublequote}term{\isachardoublequote}\ {\isacharequal}\ Var\ {\isacharprime}v\ {\isacharbar}\ App\ {\isacharprime}f\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}term\ list{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -44,7 +56,7 @@ Let us define a substitution function on terms. Because terms involve term lists, we need to define two substitution functions simultaneously:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\isanewline subst\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}v{\isasymRightarrow}{\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}term{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}term\ \ \ \ \ \ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}term{\isachardoublequote}\isanewline substs{\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}v{\isasymRightarrow}{\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}term{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}term\ list\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}term\ list{\isachardoublequote}\isanewline @@ -56,7 +68,7 @@ \ \ {\isachardoublequote}subst\ s\ {\isacharparenleft}App\ f\ ts{\isacharparenright}\ {\isacharequal}\ App\ f\ {\isacharparenleft}substs\ s\ ts{\isacharparenright}{\isachardoublequote}\isanewline \isanewline \ \ {\isachardoublequote}substs\ s\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}\isanewline -\ \ {\isachardoublequote}substs\ s\ {\isacharparenleft}t\ {\isacharhash}\ ts{\isacharparenright}\ {\isacharequal}\ subst\ s\ t\ {\isacharhash}\ substs\ s\ ts{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}substs\ s\ {\isacharparenleft}t\ {\isacharhash}\ ts{\isacharparenright}\ {\isacharequal}\ subst\ s\ t\ {\isacharhash}\ substs\ s\ ts{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -69,13 +81,26 @@ the fact that the identity substitution does not change a term needs to be strengthened and proved as follows:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ subst{\isacharunderscore}id{\isacharcolon}\ {\isachardoublequote}subst\ \ Var\ t\ \ {\isacharequal}\ {\isacharparenleft}t\ {\isacharcolon}{\isacharcolon}{\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}term{\isacharparenright}\ \ {\isasymand}\isanewline \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ substs\ Var\ ts\ {\isacharequal}\ {\isacharparenleft}ts{\isacharcolon}{\isacharcolon}{\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}term\ list{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ t\ \isakeyword{and}\ ts{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -113,33 +138,73 @@ insists on the conjunctive format. Fortunately, we can easily \emph{prove} that the suggested equation holds:% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof \isanewline +% +\endisadelimproof \isamarkupfalse% \isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}subst\ s\ {\isacharparenleft}App\ f\ ts{\isacharparenright}\ {\isacharequal}\ App\ f\ {\isacharparenleft}map\ {\isacharparenleft}subst\ s{\isacharparenright}\ ts{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ ts{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent What is more, we can now disable the old defining equation as a simplification rule:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{declare}\ subst{\isacharunderscore}App\ {\isacharbrackleft}simp\ del{\isacharbrackright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{declare}\ subst{\isacharunderscore}App\ {\isacharbrackleft}simp\ del{\isacharbrackright}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -158,8 +223,19 @@ constructor \isa{Sum} in \S\ref{sec:datatype-mut-rec} could take a list of expressions as its argument: \isa{Sum}~\isa{{\isachardoublequote}{\isacharprime}a\ aexp\ list{\isachardoublequote}}.% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Datatype/document/unfoldnested.tex --- a/doc-src/TutorialI/Datatype/document/unfoldnested.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Datatype/document/unfoldnested.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,10 +1,34 @@ % \begin{isabellebody}% \def\isabellecontext{unfoldnested}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \isamarkupfalse% \isacommand{datatype}\ {\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}{\isachardoublequote}term{\isachardoublequote}\ {\isacharequal}\ Var\ {\isacharprime}v\ {\isacharbar}\ App\ {\isacharprime}f\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}term{\isacharunderscore}list{\isachardoublequote}\isanewline -\isakeyword{and}\ {\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}term{\isacharunderscore}list\ {\isacharequal}\ Nil\ {\isacharbar}\ Cons\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}term{\isachardoublequote}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}term{\isacharunderscore}list{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% +\isakeyword{and}\ {\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}term{\isacharunderscore}list\ {\isacharequal}\ Nil\ {\isacharbar}\ Cons\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}term{\isachardoublequote}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}v{\isacharcomma}{\isacharprime}f{\isacharparenright}term{\isacharunderscore}list{\isachardoublequote}% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Documents/document/Documents.tex --- a/doc-src/TutorialI/Documents/document/Documents.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Documents/document/Documents.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Documents}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsection{Concrete Syntax \label{sec:concrete-syntax}% } @@ -41,10 +54,10 @@ case of mixfixes. The following example of the exclusive-or operation on boolean values illustrates typical infix declarations.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{constdefs}\isanewline \ \ xor\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}bool\ {\isasymRightarrow}\ bool\ {\isasymRightarrow}\ bool{\isachardoublequote}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isacharbrackleft}{\isacharplus}{\isacharbrackright}{\isachardoublequote}\ {\isadigit{6}}{\isadigit{0}}{\isacharparenright}\isanewline -\ \ {\isachardoublequote}A\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ B\ {\isasymequiv}\ {\isacharparenleft}A\ {\isasymand}\ {\isasymnot}\ B{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymnot}\ A\ {\isasymand}\ B{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}A\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ B\ {\isasymequiv}\ {\isacharparenleft}A\ {\isasymand}\ {\isasymnot}\ B{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymnot}\ A\ {\isasymand}\ B{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent Now \isa{xor\ A\ B} and \isa{A\ {\isacharbrackleft}{\isacharplus}{\isacharbrackright}\ B} refer to the @@ -142,13 +155,23 @@ \medskip Replacing our definition of \isa{xor} by the following specifies an Isabelle symbol for the new operator:% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimML +% +\endisadelimML +% +\isatagML +% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML \isamarkupfalse% \isacommand{constdefs}\isanewline \ \ xor\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}bool\ {\isasymRightarrow}\ bool\ {\isasymRightarrow}\ bool{\isachardoublequote}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isasymoplus}{\isachardoublequote}\ {\isadigit{6}}{\isadigit{0}}{\isacharparenright}\isanewline -\ \ {\isachardoublequote}A\ {\isasymoplus}\ B\ {\isasymequiv}\ {\isacharparenleft}A\ {\isasymand}\ {\isasymnot}\ B{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymnot}\ A\ {\isasymand}\ B{\isacharparenright}{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% +\ \ {\isachardoublequote}A\ {\isasymoplus}\ B\ {\isasymequiv}\ {\isacharparenleft}A\ {\isasymand}\ {\isasymnot}\ B{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymnot}\ A\ {\isasymand}\ B{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent The X-Symbol package within Proof~General provides several @@ -162,8 +185,19 @@ Proof~General's X-Symbol mode or {\LaTeX} output is active. Now consider the following hybrid declaration of \isa{xor}:% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimML +% +\endisadelimML +% +\isatagML +% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML \isamarkupfalse% \isacommand{constdefs}\isanewline \ \ xor\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}bool\ {\isasymRightarrow}\ bool\ {\isasymRightarrow}\ bool{\isachardoublequote}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isacharbrackleft}{\isacharplus}{\isacharbrackright}{\isasymignore}{\isachardoublequote}\ {\isadigit{6}}{\isadigit{0}}{\isacharparenright}\isanewline @@ -171,8 +205,7 @@ \isanewline \isamarkupfalse% \isacommand{syntax}\ {\isacharparenleft}xsymbols{\isacharparenright}\isanewline -\ \ xor\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}bool\ {\isasymRightarrow}\ bool\ {\isasymRightarrow}\ bool{\isachardoublequote}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isasymoplus}{\isasymignore}{\isachardoublequote}\ {\isadigit{6}}{\isadigit{0}}{\isacharparenright}\isamarkupfalse% -\isamarkupfalse% +\ \ xor\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}bool\ {\isasymRightarrow}\ bool\ {\isasymRightarrow}\ bool{\isachardoublequote}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isasymoplus}{\isasymignore}{\isachardoublequote}\ {\isadigit{6}}{\isadigit{0}}{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptext}% The \commdx{syntax} command introduced here acts like @@ -199,12 +232,12 @@ priorities --- just some literal syntax. The following example associates common symbols with the constructors of a datatype.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{datatype}\ currency\ {\isacharequal}\isanewline \ \ \ \ Euro\ nat\ \ \ \ {\isacharparenleft}{\isachardoublequote}{\isasymeuro}{\isachardoublequote}{\isacharparenright}\isanewline \ \ {\isacharbar}\ Pounds\ nat\ \ {\isacharparenleft}{\isachardoublequote}{\isasympounds}{\isachardoublequote}{\isacharparenright}\isanewline \ \ {\isacharbar}\ Yen\ nat\ \ \ \ \ {\isacharparenleft}{\isachardoublequote}{\isasymyen}{\isachardoublequote}{\isacharparenright}\isanewline -\ \ {\isacharbar}\ Dollar\ nat\ \ {\isacharparenleft}{\isachardoublequote}{\isachardollar}{\isachardoublequote}{\isacharparenright}\isamarkupfalse% +\ \ {\isacharbar}\ Dollar\ nat\ \ {\isacharparenleft}{\isachardoublequote}{\isachardollar}{\isachardoublequote}{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptext}% \noindent Here the mixfix annotations on the rightmost column happen @@ -243,7 +276,7 @@ relational notation for membership in a set of pair, replacing \ \isa{{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}\ {\isasymin}\ sim} by \isa{x\ {\isasymapprox}\ y}.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\isanewline \ \ sim\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set{\isachardoublequote}\isanewline \isanewline @@ -252,7 +285,7 @@ \ \ {\isachardoublequote}{\isacharunderscore}sim{\isachardoublequote}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ bool{\isachardoublequote}\ \ \ \ {\isacharparenleft}\isakeyword{infix}\ {\isachardoublequote}{\isasymapprox}{\isachardoublequote}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{translations}\isanewline -\ \ {\isachardoublequote}x\ {\isasymapprox}\ y{\isachardoublequote}\ {\isasymrightleftharpoons}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}\ {\isasymin}\ sim{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}x\ {\isasymapprox}\ y{\isachardoublequote}\ {\isasymrightleftharpoons}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}\ {\isasymin}\ sim{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent Here the name of the dummy constant \isa{{\isacharunderscore}sim} does @@ -267,10 +300,10 @@ as \isa{{\isasymnoteq}} for negated equalities. The following declaration stems from Isabelle/HOL itself:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{syntax}\ {\isachardoublequote}{\isacharunderscore}not{\isacharunderscore}equal{\isachardoublequote}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ bool{\isachardoublequote}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isasymnoteq}{\isasymignore}{\isachardoublequote}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{translations}\ {\isachardoublequote}x\ {\isasymnoteq}{\isasymignore}\ y{\isachardoublequote}\ {\isasymrightleftharpoons}\ {\isachardoublequote}{\isasymnot}\ {\isacharparenleft}x\ {\isacharequal}\ y{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\isacommand{translations}\ {\isachardoublequote}x\ {\isasymnoteq}{\isasymignore}\ y{\isachardoublequote}\ {\isasymrightleftharpoons}\ {\isachardoublequote}{\isasymnot}\ {\isacharparenleft}x\ {\isacharequal}\ y{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent Normally one would introduce derived concepts like this @@ -309,17 +342,17 @@ Here is an example to illustrate the idea of Isabelle document preparation.% \end{isamarkuptext}% -\isamarkuptrue% % \begin{quotation} +\isamarkuptrue% % \begin{isamarkuptext}% The following datatype definition of \isa{{\isacharprime}a\ bintree} models binary trees with nodes being decorated by elements of type \isa{{\isacharprime}a}.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{datatype}\ {\isacharprime}a\ bintree\ {\isacharequal}\isanewline -\ \ \ \ \ Leaf\ {\isacharbar}\ Branch\ {\isacharprime}a\ \ {\isachardoublequote}{\isacharprime}a\ bintree{\isachardoublequote}\ \ {\isachardoublequote}{\isacharprime}a\ bintree{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ Leaf\ {\isacharbar}\ Branch\ {\isacharprime}a\ \ {\isachardoublequote}{\isacharprime}a\ bintree{\isachardoublequote}\ \ {\isachardoublequote}{\isacharprime}a\ bintree{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent The datatype induction rule generated here is of the form @@ -330,9 +363,9 @@ \isaindent{\ }{\isasymLongrightarrow}\ P\ bintree% \end{isabelle}% \end{isamarkuptext}% -\isamarkuptrue% % \end{quotation} +\isamarkuptrue% % \begin{isamarkuptext}% \noindent The above document output has been produced as follows: @@ -576,7 +609,7 @@ marginal comments may be given at the same time. Here is a simple example:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}A\ {\isacharminus}{\isacharminus}{\isachargreater}\ A{\isachardoublequote}\isanewline \ \ % \isamarkupcmt{a triviality of propositional logic% @@ -586,11 +619,24 @@ \isamarkupcmt{(should not really bother)% } \isanewline -\ \ \isamarkupfalse% +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isamarkupfalse% \isacommand{by}\ {\isacharparenleft}rule\ impI{\isacharparenright}\ % \isamarkupcmt{implicit assumption step involved here% } -\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent The above output has been produced as follows: @@ -812,10 +858,23 @@ than they really were. For example, this ``fully automatic'' proof is actually a fake:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}x\ {\isasymnoteq}\ {\isacharparenleft}{\isadigit{0}}{\isacharcolon}{\isacharcolon}int{\isacharparenright}\ {\isasymLongrightarrow}\ {\isadigit{0}}\ {\isacharless}\ x\ {\isacharasterisk}\ x{\isachardoublequote}\isanewline -\ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}auto{\isacharparenright}\isamarkupfalse% +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}auto{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Here the real source of the proof has been as follows: @@ -846,8 +905,19 @@ close parentheses need to be inserted carefully; it is easy to hide the wrong parts, especially after rearranging the theory text.% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Ifexpr/document/Ifexpr.tex --- a/doc-src/TutorialI/Ifexpr/document/Ifexpr.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Ifexpr/document/Ifexpr.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Ifexpr}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsubsection{Case Study: Boolean Expressions% } @@ -24,9 +37,9 @@ constants by negation and conjunction. The following datatype serves exactly that purpose:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{datatype}\ boolex\ {\isacharequal}\ Const\ bool\ {\isacharbar}\ Var\ nat\ {\isacharbar}\ Neg\ boolex\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ And\ boolex\ boolex\isamarkupfalse% +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ And\ boolex\ boolex\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -43,14 +56,14 @@ \emph{environment} of type \isa{nat\ {\isasymRightarrow}\ bool}, which maps variables to their values:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ value\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}boolex\ {\isasymRightarrow}\ {\isacharparenleft}nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isasymRightarrow}\ bool{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\isanewline {\isachardoublequote}value\ {\isacharparenleft}Const\ b{\isacharparenright}\ env\ {\isacharequal}\ b{\isachardoublequote}\isanewline {\isachardoublequote}value\ {\isacharparenleft}Var\ x{\isacharparenright}\ \ \ env\ {\isacharequal}\ env\ x{\isachardoublequote}\isanewline {\isachardoublequote}value\ {\isacharparenleft}Neg\ b{\isacharparenright}\ \ \ env\ {\isacharequal}\ {\isacharparenleft}{\isasymnot}\ value\ b\ env{\isacharparenright}{\isachardoublequote}\isanewline -{\isachardoublequote}value\ {\isacharparenleft}And\ b\ c{\isacharparenright}\ env\ {\isacharequal}\ {\isacharparenleft}value\ b\ env\ {\isasymand}\ value\ c\ env{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +{\isachardoublequote}value\ {\isacharparenleft}And\ b\ c{\isacharparenright}\ env\ {\isacharequal}\ {\isacharparenleft}value\ b\ env\ {\isasymand}\ value\ c\ env{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -61,21 +74,21 @@ from constants (\isa{CIF}), variables (\isa{VIF}) and conditionals (\isa{IF}):% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\ ifex\ {\isacharequal}\ CIF\ bool\ {\isacharbar}\ VIF\ nat\ {\isacharbar}\ IF\ ifex\ ifex\ ifex\isamarkupfalse% +\isamarkupfalse% +\isacommand{datatype}\ ifex\ {\isacharequal}\ CIF\ bool\ {\isacharbar}\ VIF\ nat\ {\isacharbar}\ IF\ ifex\ ifex\ ifex\isamarkuptrue% % \begin{isamarkuptext}% \noindent The evaluation of If-expressions proceeds as for \isa{boolex}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ valif\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}ifex\ {\isasymRightarrow}\ {\isacharparenleft}nat\ {\isasymRightarrow}\ bool{\isacharparenright}\ {\isasymRightarrow}\ bool{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\isanewline {\isachardoublequote}valif\ {\isacharparenleft}CIF\ b{\isacharparenright}\ \ \ \ env\ {\isacharequal}\ b{\isachardoublequote}\isanewline {\isachardoublequote}valif\ {\isacharparenleft}VIF\ x{\isacharparenright}\ \ \ \ env\ {\isacharequal}\ env\ x{\isachardoublequote}\isanewline {\isachardoublequote}valif\ {\isacharparenleft}IF\ b\ t\ e{\isacharparenright}\ env\ {\isacharequal}\ {\isacharparenleft}if\ valif\ b\ env\ then\ valif\ t\ env\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ else\ valif\ e\ env{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ else\ valif\ e\ env{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \subsubsection{Converting Boolean and If-Expressions} @@ -84,33 +97,46 @@ formulae, whereas \isa{ifex} is designed for efficiency. It is easy to translate from \isa{boolex} into \isa{ifex}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ bool{\isadigit{2}}if\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}boolex\ {\isasymRightarrow}\ ifex{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\isanewline {\isachardoublequote}bool{\isadigit{2}}if\ {\isacharparenleft}Const\ b{\isacharparenright}\ {\isacharequal}\ CIF\ b{\isachardoublequote}\isanewline {\isachardoublequote}bool{\isadigit{2}}if\ {\isacharparenleft}Var\ x{\isacharparenright}\ \ \ {\isacharequal}\ VIF\ x{\isachardoublequote}\isanewline {\isachardoublequote}bool{\isadigit{2}}if\ {\isacharparenleft}Neg\ b{\isacharparenright}\ \ \ {\isacharequal}\ IF\ {\isacharparenleft}bool{\isadigit{2}}if\ b{\isacharparenright}\ {\isacharparenleft}CIF\ False{\isacharparenright}\ {\isacharparenleft}CIF\ True{\isacharparenright}{\isachardoublequote}\isanewline -{\isachardoublequote}bool{\isadigit{2}}if\ {\isacharparenleft}And\ b\ c{\isacharparenright}\ {\isacharequal}\ IF\ {\isacharparenleft}bool{\isadigit{2}}if\ b{\isacharparenright}\ {\isacharparenleft}bool{\isadigit{2}}if\ c{\isacharparenright}\ {\isacharparenleft}CIF\ False{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +{\isachardoublequote}bool{\isadigit{2}}if\ {\isacharparenleft}And\ b\ c{\isacharparenright}\ {\isacharequal}\ IF\ {\isacharparenleft}bool{\isadigit{2}}if\ b{\isacharparenright}\ {\isacharparenleft}bool{\isadigit{2}}if\ c{\isacharparenright}\ {\isacharparenleft}CIF\ False{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent At last, we have something we can verify: that \isa{bool{\isadigit{2}}if} preserves the value of its argument:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}valif\ {\isacharparenleft}bool{\isadigit{2}}if\ b{\isacharparenright}\ env\ {\isacharequal}\ value\ b\ env{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{lemma}\ {\isachardoublequote}valif\ {\isacharparenleft}bool{\isadigit{2}}if\ b{\isacharparenright}\ env\ {\isacharequal}\ value\ b\ env{\isachardoublequote}\isamarkupfalse% % \begin{isamarkuptxt}% \noindent The proof is canonical:% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ b{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -124,7 +150,7 @@ \isa{IF\ b\ {\isacharparenleft}IF\ x\ z\ u{\isacharparenright}\ {\isacharparenleft}IF\ y\ z\ u{\isacharparenright}}, which has the same value. The following primitive recursive functions perform this task:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ normif\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}ifex\ {\isasymRightarrow}\ ifex\ {\isasymRightarrow}\ ifex\ {\isasymRightarrow}\ ifex{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\isanewline @@ -138,7 +164,7 @@ \isacommand{primrec}\isanewline {\isachardoublequote}norm\ {\isacharparenleft}CIF\ b{\isacharparenright}\ \ \ \ {\isacharequal}\ CIF\ b{\isachardoublequote}\isanewline {\isachardoublequote}norm\ {\isacharparenleft}VIF\ x{\isacharparenright}\ \ \ \ {\isacharequal}\ VIF\ x{\isachardoublequote}\isanewline -{\isachardoublequote}norm\ {\isacharparenleft}IF\ b\ t\ e{\isacharparenright}\ {\isacharequal}\ normif\ b\ {\isacharparenleft}norm\ t{\isacharparenright}\ {\isacharparenleft}norm\ e{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +{\isachardoublequote}norm\ {\isacharparenleft}IF\ b\ t\ e{\isacharparenright}\ {\isacharequal}\ normif\ b\ {\isacharparenleft}norm\ t{\isacharparenright}\ {\isacharparenleft}norm\ e{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -146,23 +172,56 @@ intuitive understanding. Fortunately, Isabelle can help us to verify that the transformation preserves the value of the expression:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{theorem}\ {\isachardoublequote}valif\ {\isacharparenleft}norm\ b{\isacharparenright}\ env\ {\isacharequal}\ valif\ b\ env{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{theorem}\ {\isachardoublequote}valif\ {\isacharparenleft}norm\ b{\isacharparenright}\ env\ {\isacharequal}\ valif\ b\ env{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% % \begin{isamarkuptext}% \noindent The proof is canonical, provided we first show the following simplification lemma, which also helps to understand what \isa{normif} does:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequote}{\isasymforall}t\ e{\isachardot}\ valif\ {\isacharparenleft}normif\ b\ t\ e{\isacharparenright}\ env\ {\isacharequal}\ valif\ {\isacharparenleft}IF\ b\ t\ e{\isacharparenright}\ env{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% +\ \ {\isachardoublequote}{\isasymforall}t\ e{\isachardot}\ valif\ {\isacharparenleft}normif\ b\ t\ e{\isacharparenright}\ env\ {\isacharequal}\ valif\ {\isacharparenleft}IF\ b\ t\ e{\isacharparenright}\ env{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -172,27 +231,48 @@ But how can we be sure that \isa{norm} really produces a normal form in the above sense? We define a function that tests If-expressions for normality:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ normal\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}ifex\ {\isasymRightarrow}\ bool{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\isanewline {\isachardoublequote}normal{\isacharparenleft}CIF\ b{\isacharparenright}\ {\isacharequal}\ True{\isachardoublequote}\isanewline {\isachardoublequote}normal{\isacharparenleft}VIF\ x{\isacharparenright}\ {\isacharequal}\ True{\isachardoublequote}\isanewline {\isachardoublequote}normal{\isacharparenleft}IF\ b\ t\ e{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}normal\ t\ {\isasymand}\ normal\ e\ {\isasymand}\isanewline -\ \ \ \ \ {\isacharparenleft}case\ b\ of\ CIF\ b\ {\isasymRightarrow}\ True\ {\isacharbar}\ VIF\ x\ {\isasymRightarrow}\ True\ {\isacharbar}\ IF\ x\ y\ z\ {\isasymRightarrow}\ False{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ {\isacharparenleft}case\ b\ of\ CIF\ b\ {\isasymRightarrow}\ True\ {\isacharbar}\ VIF\ x\ {\isasymRightarrow}\ True\ {\isacharbar}\ IF\ x\ y\ z\ {\isasymRightarrow}\ False{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent Now we prove \isa{normal\ {\isacharparenleft}norm\ b{\isacharparenright}}. Of course, this requires a lemma about normality of \isa{normif}:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}{\isasymforall}t\ e{\isachardot}\ normal{\isacharparenleft}normif\ b\ t\ e{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}normal\ t\ {\isasymand}\ normal\ e{\isacharparenright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}{\isasymforall}t\ e{\isachardot}\ normal{\isacharparenleft}normif\ b\ t\ e{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}normal\ t\ {\isasymand}\ normal\ e{\isacharparenright}{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% % \begin{isamarkuptext}% \medskip @@ -211,26 +291,71 @@ \end{exercise} \index{boolean expressions example|)}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Inductive/document/AB.tex --- a/doc-src/TutorialI/Inductive/document/AB.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Inductive/document/AB.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{AB}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsection{Case Study: A Context Free Grammar% } @@ -28,27 +41,40 @@ We start by fixing the alphabet, which consists only of \isa{a}'s and~\isa{b}'s:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\ alfa\ {\isacharequal}\ a\ {\isacharbar}\ b\isamarkupfalse% +\isamarkupfalse% +\isacommand{datatype}\ alfa\ {\isacharequal}\ a\ {\isacharbar}\ b\isamarkuptrue% % \begin{isamarkuptext}% \noindent For convenience we include the following easy lemmas as simplification rules:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x\ {\isasymnoteq}\ a{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharequal}\ b{\isacharparenright}\ {\isasymand}\ {\isacharparenleft}x\ {\isasymnoteq}\ b{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharequal}\ a{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}case{\isacharunderscore}tac\ x{\isacharcomma}\ auto{\isacharparenright}\isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}case{\isacharunderscore}tac\ x{\isacharcomma}\ auto{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Words over this alphabet are of type \isa{alfa\ list}, and the three nonterminals are declared as sets of such words:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ S\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}alfa\ list\ set{\isachardoublequote}\isanewline \ \ \ \ \ \ \ A\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}alfa\ list\ set{\isachardoublequote}\isanewline -\ \ \ \ \ \ \ B\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}alfa\ list\ set{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ \ \ B\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}alfa\ list\ set{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -56,7 +82,7 @@ definition\index{inductive definition!simultaneous} of \isa{S}, \isa{A} and~\isa{B}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{inductive}\ S\ A\ B\isanewline \isakeyword{intros}\isanewline \ \ {\isachardoublequote}{\isacharbrackleft}{\isacharbrackright}\ {\isasymin}\ S{\isachardoublequote}\isanewline @@ -67,7 +93,7 @@ \ \ {\isachardoublequote}{\isasymlbrakk}\ v{\isasymin}A{\isacharsemicolon}\ w{\isasymin}A\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ b{\isacharhash}v{\isacharat}w\ {\isasymin}\ A{\isachardoublequote}\isanewline \isanewline \ \ {\isachardoublequote}w\ {\isasymin}\ S\ \ \ \ \ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ b{\isacharhash}w\ \ \ {\isasymin}\ B{\isachardoublequote}\isanewline -\ \ {\isachardoublequote}{\isasymlbrakk}\ v\ {\isasymin}\ B{\isacharsemicolon}\ w\ {\isasymin}\ B\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ a{\isacharhash}v{\isacharat}w\ {\isasymin}\ B{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}{\isasymlbrakk}\ v\ {\isasymin}\ B{\isacharsemicolon}\ w\ {\isasymin}\ B\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ a{\isacharhash}v{\isacharat}w\ {\isasymin}\ B{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -75,11 +101,17 @@ induction, so is the proof: we show at the same time that all words in \isa{A} contain one more \isa{a} than \isa{b} and all words in \isa{B} contains one more \isa{b} than \isa{a}.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ correctness{\isacharcolon}\isanewline \ \ {\isachardoublequote}{\isacharparenleft}w\ {\isasymin}\ S\ {\isasymlongrightarrow}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}a{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}b{\isacharbrackright}{\isacharparenright}\ \ \ \ \ {\isasymand}\isanewline \ \ \ {\isacharparenleft}w\ {\isasymin}\ A\ {\isasymlongrightarrow}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}a{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}b{\isacharbrackright}\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}\ {\isasymand}\isanewline -\ \ \ {\isacharparenleft}w\ {\isasymin}\ B\ {\isasymlongrightarrow}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}b{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}a{\isacharbrackright}\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ \ {\isacharparenleft}w\ {\isasymin}\ B\ {\isasymlongrightarrow}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}b{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}a{\isacharbrackright}\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -88,8 +120,15 @@ The proof itself is by rule induction and afterwards automatic:% \end{isamarkuptxt}% +\isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}rule\ S{\isacharunderscore}A{\isacharunderscore}B{\isachardot}induct{\isacharcomma}\ auto{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{by}\ {\isacharparenleft}rule\ S{\isacharunderscore}A{\isacharunderscore}B{\isachardot}induct{\isacharcomma}\ auto{\isacharparenright}\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -125,10 +164,16 @@ to prove the desired lemma twice, once as stated above and once with the roles of \isa{a}'s and \isa{b}'s interchanged.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ step{\isadigit{1}}{\isacharcolon}\ {\isachardoublequote}{\isasymforall}i\ {\isacharless}\ size\ w{\isachardot}\isanewline \ \ {\isasymbar}{\isacharparenleft}int{\isacharparenleft}size{\isacharbrackleft}x{\isasymin}take\ {\isacharparenleft}i{\isacharplus}{\isadigit{1}}{\isacharparenright}\ w{\isachardot}\ P\ x{\isacharbrackright}{\isacharparenright}{\isacharminus}int{\isacharparenleft}size{\isacharbrackleft}x{\isasymin}take\ {\isacharparenleft}i{\isacharplus}{\isadigit{1}}{\isacharparenright}\ w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharparenright}{\isacharparenright}\isanewline -\ \ \ {\isacharminus}\ {\isacharparenleft}int{\isacharparenleft}size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ P\ x{\isacharbrackright}{\isacharparenright}{\isacharminus}int{\isacharparenleft}size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharparenright}{\isacharparenright}{\isasymbar}\ {\isasymle}\ {\isadigit{1}}{\isachardoublequote}\isamarkupfalse% +\ \ \ {\isacharminus}\ {\isacharparenleft}int{\isacharparenleft}size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ P\ x{\isacharbrackright}{\isacharparenright}{\isacharminus}int{\isacharparenleft}size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharparenright}{\isacharparenright}{\isasymbar}\ {\isasymle}\ {\isadigit{1}}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -143,20 +188,33 @@ so trivial induction step. Since it is essentially just arithmetic, we do not discuss it.% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ w{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}{\isacharparenleft}auto\ simp\ add{\isacharcolon}\ abs{\isacharunderscore}if\ take{\isacharunderscore}Cons\ split{\isacharcolon}\ nat{\isachardot}split{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% Finally we come to the above-mentioned lemma about cutting in half a word with two more elements of one sort than of the other sort:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ part{\isadigit{1}}{\isacharcolon}\isanewline \ {\isachardoublequote}size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ P\ x{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharplus}{\isadigit{2}}\ {\isasymLongrightarrow}\isanewline -\ \ {\isasymexists}i{\isasymle}size\ w{\isachardot}\ size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ P\ x{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharplus}{\isadigit{1}}{\isachardoublequote}\isamarkupfalse% +\ \ {\isasymexists}i{\isasymle}size\ w{\isachardot}\ size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ P\ x{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharplus}{\isadigit{1}}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -164,10 +222,17 @@ instantiated appropriately and with its first premise disposed of by lemma \isa{step{\isadigit{1}}}:% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}insert\ nat{\isadigit{0}}{\isacharunderscore}intermed{\isacharunderscore}int{\isacharunderscore}val{\isacharbrackleft}OF\ step{\isadigit{1}}{\isacharcomma}\ of\ {\isachardoublequote}P{\isachardoublequote}\ {\isachardoublequote}w{\isachardoublequote}\ {\isachardoublequote}{\isadigit{1}}{\isachardoublequote}{\isacharbrackright}{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{by}\ force\isamarkupfalse% +\isacommand{by}\ force% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -175,14 +240,27 @@ Lemma \isa{part{\isadigit{1}}} tells us only about the prefix \isa{take\ i\ w}. An easy lemma deals with the suffix \isa{drop\ i\ w}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ part{\isadigit{2}}{\isacharcolon}\isanewline \ \ {\isachardoublequote}{\isasymlbrakk}size{\isacharbrackleft}x{\isasymin}take\ i\ w\ {\isacharat}\ drop\ i\ w{\isachardot}\ P\ x{\isacharbrackright}\ {\isacharequal}\isanewline \ \ \ \ size{\isacharbrackleft}x{\isasymin}take\ i\ w\ {\isacharat}\ drop\ i\ w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharplus}{\isadigit{2}}{\isacharsemicolon}\isanewline \ \ \ \ size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ P\ x{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharplus}{\isadigit{1}}{\isasymrbrakk}\isanewline \ \ \ {\isasymLongrightarrow}\ size{\isacharbrackleft}x{\isasymin}drop\ i\ w{\isachardot}\ P\ x{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}drop\ i\ w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharplus}{\isadigit{1}}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}{\isacharparenleft}simp\ del{\isacharcolon}\ append{\isacharunderscore}take{\isacharunderscore}drop{\isacharunderscore}id{\isacharparenright}\isamarkupfalse% +\isacommand{by}{\isacharparenleft}simp\ del{\isacharcolon}\ append{\isacharunderscore}take{\isacharunderscore}drop{\isacharunderscore}id{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -199,8 +277,8 @@ To dispose of trivial cases automatically, the rules of the inductive definition are declared simplification rules:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{declare}\ S{\isacharunderscore}A{\isacharunderscore}B{\isachardot}intros{\isacharbrackleft}simp{\isacharbrackright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{declare}\ S{\isacharunderscore}A{\isacharunderscore}B{\isachardot}intros{\isacharbrackleft}simp{\isacharbrackright}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -210,11 +288,17 @@ \isa{a}'s and \isa{b}'s, then it is in \isa{S}, and similarly for \isa{A} and \isa{B}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{theorem}\ completeness{\isacharcolon}\isanewline \ \ {\isachardoublequote}{\isacharparenleft}size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}a{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}b{\isacharbrackright}\ \ \ \ \ {\isasymlongrightarrow}\ w\ {\isasymin}\ S{\isacharparenright}\ {\isasymand}\isanewline \ \ \ {\isacharparenleft}size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}a{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}b{\isacharbrackright}\ {\isacharplus}\ {\isadigit{1}}\ {\isasymlongrightarrow}\ w\ {\isasymin}\ A{\isacharparenright}\ {\isasymand}\isanewline -\ \ \ {\isacharparenleft}size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}b{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}a{\isacharbrackright}\ {\isacharplus}\ {\isadigit{1}}\ {\isasymlongrightarrow}\ w\ {\isasymin}\ B{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ \ {\isacharparenleft}size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}b{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}a{\isacharbrackright}\ {\isacharplus}\ {\isadigit{1}}\ {\isasymlongrightarrow}\ w\ {\isasymin}\ B{\isacharparenright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -223,9 +307,8 @@ merely appending a single letter at the front. Hence we induct on the length of \isa{w}, using the induction rule \isa{length{\isacharunderscore}induct}:% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ w\ rule{\isacharcolon}\ length{\isacharunderscore}induct{\isacharparenright}\isamarkupfalse% \isamarkupfalse% +\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ w\ rule{\isacharcolon}\ length{\isacharunderscore}induct{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -237,11 +320,10 @@ The proof continues with a case distinction on \isa{w}, on whether \isa{w} is empty or not.% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}case{\isacharunderscore}tac\ w{\isacharparenright}\isanewline \ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp{\isacharunderscore}all{\isacharparenright}\isamarkupfalse% -\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}simp{\isacharunderscore}all{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -256,14 +338,14 @@ After breaking the conjunction up into two cases, we can apply \isa{part{\isadigit{1}}} to the assumption that \isa{w} contains two more \isa{a}'s than \isa{b}'s.% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}rule\ conjI{\isacharparenright}\isanewline \ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}clarify{\isacharparenright}\isanewline \ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}frule\ part{\isadigit{1}}{\isacharbrackleft}of\ {\isachardoublequote}{\isasymlambda}x{\isachardot}\ x{\isacharequal}a{\isachardoublequote}{\isacharcomma}\ simplified{\isacharbrackright}{\isacharparenright}\isanewline \ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}clarify{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}clarify{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -276,18 +358,18 @@ \ \ \ \ \ length\ {\isacharbrackleft}x{\isasymin}drop\ i\ v\ {\isachardot}\ x\ {\isacharequal}\ a{\isacharbrackright}\ {\isacharequal}\ length\ {\isacharbrackleft}x{\isasymin}drop\ i\ v\ {\isachardot}\ x\ {\isacharequal}\ b{\isacharbrackright}\ {\isacharplus}\ {\isadigit{1}}% \end{isabelle}% \end{isamarkuptxt}% -\ \isamarkuptrue% +\ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}drule\ part{\isadigit{2}}{\isacharbrackleft}of\ {\isachardoublequote}{\isasymlambda}x{\isachardot}\ x{\isacharequal}a{\isachardoublequote}{\isacharcomma}\ simplified{\isacharbrackright}{\isacharparenright}\isanewline \ \ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}assumption{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}assumption{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent Now it is time to decompose \isa{v} in the conclusion \isa{b\ {\isacharhash}\ v\ {\isasymin}\ A} into \isa{take\ i\ v\ {\isacharat}\ drop\ i\ v},% \end{isamarkuptxt}% -\ \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}rule{\isacharunderscore}tac\ n{\isadigit{1}}{\isacharequal}i\ \isakeyword{and}\ t{\isacharequal}v\ \isakeyword{in}\ subst{\isacharbrackleft}OF\ append{\isacharunderscore}take{\isacharunderscore}drop{\isacharunderscore}id{\isacharbrackright}{\isacharparenright}\isamarkupfalse% +\ \isamarkupfalse% +\isacommand{apply}{\isacharparenleft}rule{\isacharunderscore}tac\ n{\isadigit{1}}{\isacharequal}i\ \isakeyword{and}\ t{\isacharequal}v\ \isakeyword{in}\ subst{\isacharbrackleft}OF\ append{\isacharunderscore}take{\isacharunderscore}drop{\isacharunderscore}id{\isacharbrackright}{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -296,21 +378,21 @@ after which the appropriate rule of the grammar reduces the goal to the two subgoals \isa{take\ i\ v\ {\isasymin}\ A} and \isa{drop\ i\ v\ {\isasymin}\ A}:% \end{isamarkuptxt}% -\ \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}rule\ S{\isacharunderscore}A{\isacharunderscore}B{\isachardot}intros{\isacharparenright}\isamarkupfalse% +\ \isamarkupfalse% +\isacommand{apply}{\isacharparenleft}rule\ S{\isacharunderscore}A{\isacharunderscore}B{\isachardot}intros{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% Both subgoals follow from the induction hypothesis because both \isa{take\ i\ v} and \isa{drop\ i\ v} are shorter than \isa{w}:% \end{isamarkuptxt}% -\ \ \isamarkuptrue% +\ \ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}force\ simp\ add{\isacharcolon}\ min{\isacharunderscore}less{\isacharunderscore}iff{\isacharunderscore}disj{\isacharparenright}\isanewline \ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}force\ split\ add{\isacharcolon}\ nat{\isacharunderscore}diff{\isacharunderscore}split{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}force\ split\ add{\isacharcolon}\ nat{\isacharunderscore}diff{\isacharunderscore}split{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% The case \isa{w\ {\isacharequal}\ b\ {\isacharhash}\ v} is proved analogously:% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}clarify{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}{\isacharparenleft}frule\ part{\isadigit{1}}{\isacharbrackleft}of\ {\isachardoublequote}{\isasymlambda}x{\isachardot}\ x{\isacharequal}b{\isachardoublequote}{\isacharcomma}\ simplified{\isacharbrackright}{\isacharparenright}\isanewline @@ -327,7 +409,14 @@ \ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}force\ simp\ add{\isacharcolon}\ min{\isacharunderscore}less{\isacharunderscore}iff{\isacharunderscore}disj{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{by}{\isacharparenleft}force\ simp\ add{\isacharcolon}\ min{\isacharunderscore}less{\isacharunderscore}iff{\isacharunderscore}disj\ split\ add{\isacharcolon}\ nat{\isacharunderscore}diff{\isacharunderscore}split{\isacharparenright}\isamarkupfalse% +\isacommand{by}{\isacharparenleft}force\ simp\ add{\isacharcolon}\ min{\isacharunderscore}less{\isacharunderscore}iff{\isacharunderscore}disj\ split\ add{\isacharcolon}\ nat{\isacharunderscore}diff{\isacharunderscore}split{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% We conclude this section with a comparison of our proof with @@ -354,8 +443,19 @@ are scrutinized formally.% \index{grammars!defining inductively|)}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Inductive/document/Advanced.tex --- a/doc-src/TutorialI/Inductive/document/Advanced.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Inductive/document/Advanced.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,8 +1,22 @@ % \begin{isabellebody}% \def\isabellecontext{Advanced}% +% +\isadelimtheory \isanewline -\isacommand{theory}\ Advanced\ \isakeyword{imports}\ Even\ \isakeyword{begin}\isanewline +% +\endisadelimtheory +% +\isatagtheory +\isamarkupfalse% +\isacommand{theory}\ Advanced\ \isakeyword{imports}\ Even\ \isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +\isanewline +% +\endisadelimtheory \isanewline \isanewline \isamarkupfalse% @@ -21,10 +35,16 @@ \isanewline \isamarkupfalse% \isacommand{lemma}\ gterms{\isacharunderscore}mono{\isacharcolon}\ {\isachardoublequote}F{\isasymsubseteq}G\ {\isasymLongrightarrow}\ gterms\ F\ {\isasymsubseteq}\ gterms\ G{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}\ clarify\isanewline \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}erule\ gterms{\isachardot}induct{\isacharparenright}\isamarkupfalse% +\isacommand{apply}\ {\isacharparenleft}erule\ gterms{\isachardot}induct{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -33,10 +53,17 @@ \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymLongrightarrow}\ Apply\ f\ args\ {\isasymin}\ gterms\ G% \end{isabelle}% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}\ blast\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \begin{isabelle}% @@ -49,12 +76,12 @@ the two forms that Markus has made available. First the one for binding the result to a name% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{inductive{\isacharunderscore}cases}\ Suc{\isacharunderscore}Suc{\isacharunderscore}cases\ {\isacharbrackleft}elim{\isacharbang}{\isacharbrackright}{\isacharcolon}\isanewline \ \ {\isachardoublequote}Suc{\isacharparenleft}Suc\ n{\isacharparenright}\ {\isasymin}\ even{\isachardoublequote}\isanewline \isanewline \isamarkupfalse% -\isacommand{thm}\ Suc{\isacharunderscore}Suc{\isacharunderscore}cases\isamarkupfalse% +\isacommand{thm}\ Suc{\isacharunderscore}Suc{\isacharunderscore}cases\isamarkuptrue% % \begin{isamarkuptext}% \begin{isabelle}% @@ -64,15 +91,28 @@ and now the one for local usage:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}Suc{\isacharparenleft}Suc\ n{\isacharparenright}\ {\isasymin}\ even\ {\isasymLongrightarrow}\ P\ n{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}\ {\isacharparenleft}ind{\isacharunderscore}cases\ {\isachardoublequote}Suc{\isacharparenleft}Suc\ n{\isacharparenright}\ {\isasymin}\ even{\isachardoublequote}{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{oops}\isanewline +\isacommand{oops}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof \isanewline \isamarkupfalse% -\isacommand{inductive{\isacharunderscore}cases}\ gterm{\isacharunderscore}Apply{\isacharunderscore}elim\ {\isacharbrackleft}elim{\isacharbang}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}Apply\ f\ args\ {\isasymin}\ gterms\ F{\isachardoublequote}\isamarkupfalse% +\isacommand{inductive{\isacharunderscore}cases}\ gterm{\isacharunderscore}Apply{\isacharunderscore}elim\ {\isacharbrackleft}elim{\isacharbang}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}Apply\ f\ args\ {\isasymin}\ gterms\ F{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% this is what we get: @@ -82,11 +122,17 @@ \end{isabelle} \rulename{gterm_Apply_elim}% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ gterms{\isacharunderscore}IntI\ {\isacharbrackleft}rule{\isacharunderscore}format{\isacharcomma}\ intro{\isacharbang}{\isacharbrackright}{\isacharcolon}\isanewline \ \ \ \ \ {\isachardoublequote}t\ {\isasymin}\ gterms\ F\ {\isasymLongrightarrow}\ t\ {\isasymin}\ gterms\ G\ {\isasymlongrightarrow}\ t\ {\isasymin}\ gterms\ {\isacharparenleft}F{\isasyminter}G{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}erule\ gterms{\isachardot}induct{\isacharparenright}\isamarkupfalse% +\isacommand{apply}\ {\isacharparenleft}erule\ gterms{\isachardot}induct{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -98,10 +144,17 @@ \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ {\isasymLongrightarrow}\ }Apply\ f\ args\ {\isasymin}\ gterms\ {\isacharparenleft}F\ {\isasyminter}\ G{\isacharparenright}% \end{isabelle}% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}\ blast\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \begin{isabelle}% @@ -109,16 +162,29 @@ \end{isabelle} \rulename{mono_Int}% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ gterms{\isacharunderscore}Int{\isacharunderscore}eq\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\isanewline \ \ \ \ \ {\isachardoublequote}gterms\ {\isacharparenleft}F{\isasyminter}G{\isacharparenright}\ {\isacharequal}\ gterms\ F\ {\isasyminter}\ gterms\ G{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}blast\ intro{\isacharbang}{\isacharcolon}\ mono{\isacharunderscore}Int\ monoI\ gterms{\isacharunderscore}mono{\isacharparenright}\isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}blast\ intro{\isacharbang}{\isacharcolon}\ mono{\isacharunderscore}Int\ monoI\ gterms{\isacharunderscore}mono{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% the following declaration isn't actually used% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ integer{\isacharunderscore}arity\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}integer{\isacharunderscore}op\ {\isasymRightarrow}\ nat{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\isanewline @@ -148,8 +214,14 @@ \isanewline \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}well{\isacharunderscore}formed{\isacharunderscore}gterm\ arity\ {\isasymsubseteq}\ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ arity{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}\ clarify\isamarkupfalse% +\isacommand{apply}\ clarify\isamarkuptrue% % \begin{isamarkuptxt}% The situation after clarify @@ -158,8 +230,8 @@ \isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ }x\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ arity% \end{isabelle}% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\ {\isacharparenleft}erule\ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isachardot}induct{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}\ {\isacharparenleft}erule\ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isachardot}induct{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% note the induction hypothesis! @@ -172,17 +244,30 @@ \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymLongrightarrow}\ Apply\ f\ args\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ arity% \end{isabelle}% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}\ auto\isanewline \isamarkupfalse% -\isacommand{done}\isanewline +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof \isanewline \isanewline \isanewline \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}\ arity\ {\isasymsubseteq}\ well{\isacharunderscore}formed{\isacharunderscore}gterm\ arity{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}\ clarify\isamarkupfalse% +\isacommand{apply}\ clarify\isamarkuptrue% % \begin{isamarkuptxt}% The situation after clarify @@ -191,8 +276,8 @@ \isaindent{\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ }x\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm\ arity% \end{isabelle}% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\ {\isacharparenleft}erule\ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}{\isachardot}induct{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}\ {\isacharparenleft}erule\ well{\isacharunderscore}formed{\isacharunderscore}gterm{\isacharprime}{\isachardot}induct{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% note the induction hypothesis! @@ -206,10 +291,17 @@ \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymLongrightarrow}\ Apply\ f\ args\ {\isasymin}\ well{\isacharunderscore}formed{\isacharunderscore}gterm\ arity% \end{isabelle}% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}\ auto\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \begin{isabelle}% @@ -221,7 +313,7 @@ \begin{isamarkuptext}% the rest isn't used: too complicated. OK for an exercise though.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ integer{\isacharunderscore}signature\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}integer{\isacharunderscore}op\ {\isacharasterisk}\ {\isacharparenleft}unit\ list\ {\isacharasterisk}\ unit{\isacharparenright}{\isacharparenright}\ set{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{inductive}\ {\isachardoublequote}integer{\isacharunderscore}signature{\isachardoublequote}\isanewline @@ -257,6 +349,12 @@ \isanewline \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}well{\isacharunderscore}typed{\isacharunderscore}gterm\ sig\ {\isasymsubseteq}\ well{\isacharunderscore}typed{\isacharunderscore}gterm{\isacharprime}\ sig{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}\ clarify\isanewline \isamarkupfalse% @@ -264,10 +362,23 @@ \isamarkupfalse% \isacommand{apply}\ auto\isanewline \isamarkupfalse% -\isacommand{done}\isanewline +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof \isanewline \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}well{\isacharunderscore}typed{\isacharunderscore}gterm{\isacharprime}\ sig\ {\isasymsubseteq}\ well{\isacharunderscore}typed{\isacharunderscore}gterm\ sig{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}\ clarify\isanewline \isamarkupfalse% @@ -275,13 +386,32 @@ \isamarkupfalse% \isacommand{apply}\ auto\isanewline \isamarkupfalse% -\isacommand{done}\isanewline +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof \isanewline +% +\endisadelimproof \isanewline +% +\isadelimtheory +\isanewline +% +\endisadelimtheory +% +\isatagtheory \isamarkupfalse% -\isacommand{end}\isanewline +\isacommand{end}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory \isanewline -\isamarkupfalse% +% +\endisadelimtheory +\isanewline \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Inductive/document/Even.tex --- a/doc-src/TutorialI/Inductive/document/Even.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Inductive/document/Even.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,8 +1,22 @@ % \begin{isabellebody}% \def\isabellecontext{Even}% +% +\isadelimtheory \isanewline -\isacommand{theory}\ Even\ \isakeyword{imports}\ Main\ \isakeyword{begin}\isanewline +% +\endisadelimtheory +% +\isatagtheory +\isamarkupfalse% +\isacommand{theory}\ Even\ \isakeyword{imports}\ Main\ \isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +\isanewline +% +\endisadelimtheory \isanewline \isanewline \isamarkupfalse% @@ -11,7 +25,7 @@ \isacommand{inductive}\ even\isanewline \isakeyword{intros}\isanewline zero{\isacharbrackleft}intro{\isacharbang}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}{\isadigit{0}}\ {\isasymin}\ even{\isachardoublequote}\isanewline -step{\isacharbrackleft}intro{\isacharbang}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}n\ {\isasymin}\ even\ {\isasymLongrightarrow}\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}\ {\isasymin}\ even{\isachardoublequote}\isamarkupfalse% +step{\isacharbrackleft}intro{\isacharbang}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}n\ {\isasymin}\ even\ {\isasymLongrightarrow}\ {\isacharparenleft}Suc\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}\ {\isasymin}\ even{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% An inductive definition consists of introduction rules. @@ -32,10 +46,16 @@ Our first lemma states that numbers of the form $2\times k$ are even.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ two{\isacharunderscore}times{\isacharunderscore}even{\isacharbrackleft}intro{\isacharbang}{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}{\isadigit{2}}{\isacharasterisk}k\ {\isasymin}\ even{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}induct{\isacharunderscore}tac\ k{\isacharparenright}\isamarkupfalse% +\isacommand{apply}\ {\isacharparenleft}induct{\isacharunderscore}tac\ k{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% The first step is induction on the natural number \isa{k}, which leaves @@ -47,10 +67,17 @@ Here \isa{auto} simplifies both subgoals so that they match the introduction rules, which then are applied automatically.% \end{isamarkuptxt}% -\ \isamarkuptrue% +\ \isamarkupfalse% \isacommand{apply}\ auto\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% Our goal is to prove the equivalence between the traditional definition @@ -58,18 +85,37 @@ this equivalence is trivial using the lemma just proved, whose \isa{intro!} attribute ensures it will be applied automatically.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ dvd{\isacharunderscore}imp{\isacharunderscore}even{\isacharcolon}\ {\isachardoublequote}{\isadigit{2}}\ dvd\ n\ {\isasymLongrightarrow}\ n\ {\isasymin}\ even{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ dvd{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}auto\ simp\ add{\isacharcolon}\ dvd{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% our first rule induction!% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ even{\isacharunderscore}imp{\isacharunderscore}dvd{\isacharcolon}\ {\isachardoublequote}n\ {\isasymin}\ even\ {\isasymLongrightarrow}\ {\isadigit{2}}\ dvd\ n{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}erule\ even{\isachardot}induct{\isacharparenright}\isamarkupfalse% +\isacommand{apply}\ {\isacharparenleft}erule\ even{\isachardot}induct{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -77,42 +123,68 @@ \ {\isadigit{2}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isasymlbrakk}n\ {\isasymin}\ Even{\isachardot}even{\isacharsemicolon}\ {\isadigit{2}}\ dvd\ n{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isadigit{2}}\ dvd\ Suc\ {\isacharparenleft}Suc\ n{\isacharparenright}% \end{isabelle}% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ dvd{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ dvd{\isacharunderscore}def{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% \ {\isadigit{1}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ {\isasymlbrakk}n\ {\isasymin}\ Even{\isachardot}even{\isacharsemicolon}\ {\isasymexists}k{\isachardot}\ n\ {\isacharequal}\ {\isadigit{2}}\ {\isacharasterisk}\ k{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isasymexists}k{\isachardot}\ Suc\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ {\isadigit{2}}\ {\isacharasterisk}\ k% \end{isabelle}% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\ clarify\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}\ clarify\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% \ {\isadigit{1}}{\isachardot}\ {\isasymAnd}n\ k{\isachardot}\ {\isadigit{2}}\ {\isacharasterisk}\ k\ {\isasymin}\ Even{\isachardot}even\ {\isasymLongrightarrow}\ {\isasymexists}ka{\isachardot}\ Suc\ {\isacharparenleft}Suc\ {\isacharparenleft}{\isadigit{2}}\ {\isacharasterisk}\ k{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ {\isadigit{2}}\ {\isacharasterisk}\ ka% \end{isabelle}% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}\ {\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isachardoublequote}Suc\ k{\isachardoublequote}\ \isakeyword{in}\ exI{\isacharcomma}\ simp{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% no iff-attribute because we don't always want to use it% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{theorem}\ even{\isacharunderscore}iff{\isacharunderscore}dvd{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}n\ {\isasymin}\ even{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isadigit{2}}\ dvd\ n{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}blast\ intro{\isacharcolon}\ dvd{\isacharunderscore}imp{\isacharunderscore}even\ even{\isacharunderscore}imp{\isacharunderscore}dvd{\isacharparenright}\isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}blast\ intro{\isacharcolon}\ dvd{\isacharunderscore}imp{\isacharunderscore}even\ even{\isacharunderscore}imp{\isacharunderscore}dvd{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% this result ISN'T inductive...% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ Suc{\isacharunderscore}Suc{\isacharunderscore}even{\isacharunderscore}imp{\isacharunderscore}even{\isacharcolon}\ {\isachardoublequote}Suc\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isasymin}\ even\ {\isasymLongrightarrow}\ n\ {\isasymin}\ even{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}erule\ even{\isachardot}induct{\isacharparenright}\isamarkupfalse% +\isacommand{apply}\ {\isacharparenleft}erule\ even{\isachardot}induct{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -120,16 +192,29 @@ \ {\isadigit{2}}{\isachardot}\ {\isasymAnd}na{\isachardot}\ {\isasymlbrakk}na\ {\isasymin}\ Even{\isachardot}even{\isacharsemicolon}\ n\ {\isasymin}\ Even{\isachardot}even{\isasymrbrakk}\ {\isasymLongrightarrow}\ n\ {\isasymin}\ Even{\isachardot}even% \end{isabelle}% \end{isamarkuptxt}% +\isamarkupfalse% +\isacommand{oops}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{oops}\isamarkupfalse% % \begin{isamarkuptext}% ...so we need an inductive lemma...% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ even{\isacharunderscore}imp{\isacharunderscore}even{\isacharunderscore}minus{\isacharunderscore}{\isadigit{2}}{\isacharcolon}\ {\isachardoublequote}n\ {\isasymin}\ even\ {\isasymLongrightarrow}\ n\ {\isacharminus}\ {\isadigit{2}}\ {\isasymin}\ even{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}erule\ even{\isachardot}induct{\isacharparenright}\isamarkupfalse% +\isacommand{apply}\ {\isacharparenleft}erule\ even{\isachardot}induct{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -138,29 +223,74 @@ \isaindent{\ {\isadigit{2}}{\isachardot}\ {\isasymAnd}n{\isachardot}\ }{\isasymLongrightarrow}\ Suc\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharminus}\ {\isadigit{2}}\ {\isasymin}\ Even{\isachardot}even% \end{isabelle}% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}\ auto\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% ...and prove it in a separate step% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ Suc{\isacharunderscore}Suc{\isacharunderscore}even{\isacharunderscore}imp{\isacharunderscore}even{\isacharcolon}\ {\isachardoublequote}Suc\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isasymin}\ even\ {\isasymLongrightarrow}\ n\ {\isasymin}\ even{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}drule\ even{\isacharunderscore}imp{\isacharunderscore}even{\isacharunderscore}minus{\isacharunderscore}{\isadigit{2}}{\isacharcomma}\ simp{\isacharparenright}\isanewline +\isacommand{by}\ {\isacharparenleft}drule\ even{\isacharunderscore}imp{\isacharunderscore}even{\isacharunderscore}minus{\isacharunderscore}{\isadigit{2}}{\isacharcomma}\ simp{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof \isanewline \isanewline \isamarkupfalse% \isacommand{lemma}\ {\isacharbrackleft}iff{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharparenleft}Suc\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}\ {\isasymin}\ even{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}n\ {\isasymin}\ even{\isacharparenright}{\isachardoublequote}\isanewline -\isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}blast\ dest{\isacharcolon}\ Suc{\isacharunderscore}Suc{\isacharunderscore}even{\isacharunderscore}imp{\isacharunderscore}even{\isacharparenright}\isanewline -\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{end}\isanewline +\isacommand{by}\ {\isacharparenleft}blast\ dest{\isacharcolon}\ Suc{\isacharunderscore}Suc{\isacharunderscore}even{\isacharunderscore}imp{\isacharunderscore}even{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof +% +\isadelimtheory \isanewline +% +\endisadelimtheory +% +\isatagtheory \isamarkupfalse% +\isacommand{end}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +\isanewline +% +\endisadelimtheory +\isanewline \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Inductive/document/Mutual.tex --- a/doc-src/TutorialI/Inductive/document/Mutual.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Inductive/document/Mutual.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Mutual}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsubsection{Mutually Inductive Definitions% } @@ -12,7 +25,7 @@ by mutual induction. As a trivial example we consider the even and odd natural numbers:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ even\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat\ set{\isachardoublequote}\isanewline \ \ \ \ \ \ \ odd\ \ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat\ set{\isachardoublequote}\isanewline \isanewline @@ -21,7 +34,7 @@ \isakeyword{intros}\isanewline zero{\isacharcolon}\ \ {\isachardoublequote}{\isadigit{0}}\ {\isasymin}\ even{\isachardoublequote}\isanewline evenI{\isacharcolon}\ {\isachardoublequote}n\ {\isasymin}\ odd\ {\isasymLongrightarrow}\ Suc\ n\ {\isasymin}\ even{\isachardoublequote}\isanewline -oddI{\isacharcolon}\ \ {\isachardoublequote}n\ {\isasymin}\ even\ {\isasymLongrightarrow}\ Suc\ n\ {\isasymin}\ odd{\isachardoublequote}\isamarkupfalse% +oddI{\isacharcolon}\ \ {\isachardoublequote}n\ {\isasymin}\ even\ {\isasymLongrightarrow}\ Suc\ n\ {\isasymin}\ odd{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -37,8 +50,14 @@ If we want to prove that all even numbers are divisible by two, we have to generalize the statement as follows:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}m\ {\isasymin}\ even\ {\isasymlongrightarrow}\ {\isadigit{2}}\ dvd\ m{\isacharparenright}\ {\isasymand}\ {\isacharparenleft}n\ {\isasymin}\ odd\ {\isasymlongrightarrow}\ {\isadigit{2}}\ dvd\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}m\ {\isasymin}\ even\ {\isasymlongrightarrow}\ {\isadigit{2}}\ dvd\ m{\isacharparenright}\ {\isasymand}\ {\isacharparenleft}n\ {\isasymin}\ odd\ {\isasymlongrightarrow}\ {\isadigit{2}}\ dvd\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% % \begin{isamarkuptxt}% \noindent @@ -46,8 +65,8 @@ it is applied by \isa{rule} rather than \isa{erule} as for ordinary inductive definitions:% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}{\isacharparenleft}rule\ even{\isacharunderscore}odd{\isachardot}induct{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}rule\ even{\isacharunderscore}odd{\isachardot}induct{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -60,15 +79,26 @@ where the same subgoal was encountered before. We do not show the proof script.% \end{isamarkuptxt}% -\isamarkuptrue% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Inductive/document/Star.tex --- a/doc-src/TutorialI/Inductive/document/Star.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Inductive/document/Star.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Star}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsection{The Reflexive Transitive Closure% } @@ -19,13 +32,13 @@ defined as a least fixed point because inductive definitions were not yet available. But now they are:% \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 rtc{\isacharunderscore}refl{\isacharbrackleft}iff{\isacharbrackright}{\isacharcolon}\ \ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}x{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isanewline -rtc{\isacharunderscore}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% +rtc{\isacharunderscore}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 @@ -44,10 +57,23 @@ The rest of this section is devoted to proving that it is equivalent to the standard definition. We start with a simple lemma:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isacharbrackleft}intro{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}{\isacharparenleft}blast\ intro{\isacharcolon}\ rtc{\isacharunderscore}step{\isacharparenright}\isamarkupfalse% +\isacommand{by}{\isacharparenleft}blast\ intro{\isacharcolon}\ rtc{\isacharunderscore}step{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -73,10 +99,16 @@ Now we turn to the inductive proof of transitivity:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ rtc{\isacharunderscore}trans{\isacharcolon}\ {\isachardoublequote}{\isasymlbrakk}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isacharsemicolon}\ {\isacharparenleft}y{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}erule\ rtc{\isachardot}induct{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}erule\ rtc{\isachardot}induct{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -102,10 +134,22 @@ weak. Fortunately, it can easily be strengthened: transfer the additional premise \isa{{\isacharparenleft}y{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}} into the conclusion:% \end{isamarkuptxt}% -\isamarkuptrue% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkupfalse% \isacommand{lemma}\ rtc{\isacharunderscore}trans{\isacharbrackleft}rule{\isacharunderscore}format{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymLongrightarrow}\ {\isacharparenleft}y{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymlongrightarrow}\ {\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymLongrightarrow}\ {\isacharparenleft}y{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymlongrightarrow}\ {\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -120,8 +164,8 @@ \isa{{\isasymlongrightarrow}} back into \isa{{\isasymLongrightarrow}}: in the end we obtain the original statement of our lemma.% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}{\isacharparenleft}erule\ rtc{\isachardot}induct{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}erule\ rtc{\isachardot}induct{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -133,34 +177,47 @@ \isaindent{\ {\isadigit{2}}{\isachardot}\ \ \ \ }{\isasymLongrightarrow}\ {\isacharparenleft}za{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymlongrightarrow}\ {\isacharparenleft}x{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}% \end{isabelle}% \end{isamarkuptxt}% -\ \isamarkuptrue% +\ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}blast{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ rtc{\isacharunderscore}step{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% Let us now prove that \isa{r{\isacharasterisk}} is really the reflexive transitive closure of \isa{r}, i.e.\ the least reflexive and transitive relation containing \isa{r}. The latter is easily formalized% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ rtc{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}set\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}set{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{inductive}\ {\isachardoublequote}rtc{\isadigit{2}}\ r{\isachardoublequote}\isanewline \isakeyword{intros}\isanewline {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ rtc{\isadigit{2}}\ r{\isachardoublequote}\isanewline {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}x{\isacharparenright}\ {\isasymin}\ rtc{\isadigit{2}}\ r{\isachardoublequote}\isanewline -{\isachardoublequote}{\isasymlbrakk}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ rtc{\isadigit{2}}\ r{\isacharsemicolon}\ {\isacharparenleft}y{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ rtc{\isadigit{2}}\ r\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ rtc{\isadigit{2}}\ r{\isachardoublequote}\isamarkupfalse% +{\isachardoublequote}{\isasymlbrakk}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ rtc{\isadigit{2}}\ r{\isacharsemicolon}\ {\isacharparenleft}y{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ rtc{\isadigit{2}}\ r\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}z{\isacharparenright}\ {\isasymin}\ rtc{\isadigit{2}}\ r{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent and the equivalence of the two definitions is easily shown by the obvious rule inductions:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ rtc{\isadigit{2}}\ r\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}erule\ rtc{\isadigit{2}}{\isachardot}induct{\isacharparenright}\isanewline \ \ \isamarkupfalse% @@ -170,10 +227,23 @@ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ rtc{\isacharunderscore}trans{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isanewline +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof \isanewline \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ r{\isacharasterisk}\ {\isasymLongrightarrow}\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isasymin}\ rtc{\isadigit{2}}\ r{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}erule\ rtc{\isachardot}induct{\isacharparenright}\isanewline \ \isamarkupfalse% @@ -181,7 +251,14 @@ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ rtc{\isadigit{2}}{\isachardot}intros{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% So why did we start with the first definition? Because it is simpler. It @@ -204,13 +281,32 @@ in exercise~\ref{ex:converse-rtc-step}. \end{exercise}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Misc/document/AdvancedInd.tex --- a/doc-src/TutorialI/Misc/document/AdvancedInd.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Misc/document/AdvancedInd.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{AdvancedInd}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -26,10 +39,16 @@ Since \isa{hd} and \isa{last} return the first and last element of a non-empty list, this lemma looks easy to prove:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymLongrightarrow}\ hd{\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ last\ xs{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -62,10 +81,21 @@ \attrdx{rule_format} (\S\ref{sec:forward}) convert the result to the usual \isa{{\isasymLongrightarrow}} form:% \end{isamarkuptxt}% -\isamarkuptrue% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkupfalse% -\isacommand{lemma}\ hd{\isacharunderscore}rev\ {\isacharbrackleft}rule{\isacharunderscore}format{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongrightarrow}\ hd{\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ last\ xs{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% +\isacommand{lemma}\ hd{\isacharunderscore}rev\ {\isacharbrackleft}rule{\isacharunderscore}format{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymlongrightarrow}\ hd{\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ last\ xs{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -122,8 +152,14 @@ single theorem because it depends on the number of free variables in $t$ --- the notation $\overline{y}$ is merely an informal device.% \end{isamarkuptxt}% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isamarkupfalse% % \isamarkupsubsection{Beyond Structural and Recursion Induction% } @@ -149,10 +185,10 @@ As an application, we prove a property of the following function:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ f\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat\ {\isasymRightarrow}\ nat{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{axioms}\ f{\isacharunderscore}ax{\isacharcolon}\ {\isachardoublequote}f{\isacharparenleft}f{\isacharparenleft}n{\isacharparenright}{\isacharparenright}\ {\isacharless}\ f{\isacharparenleft}Suc{\isacharparenleft}n{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\isacommand{axioms}\ f{\isacharunderscore}ax{\isacharcolon}\ {\isachardoublequote}f{\isacharparenleft}f{\isacharparenleft}n{\isacharparenright}{\isacharparenright}\ {\isacharless}\ f{\isacharparenleft}Suc{\isacharparenleft}n{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \begin{warn} @@ -168,8 +204,14 @@ be proved by induction on \mbox{\isa{f\ n}}. Following the recipe outlined above, we have to phrase the proposition as follows to allow induction:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ f{\isacharunderscore}incr{\isacharunderscore}lem{\isacharcolon}\ {\isachardoublequote}{\isasymforall}i{\isachardot}\ k\ {\isacharequal}\ f\ i\ {\isasymlongrightarrow}\ i\ {\isasymle}\ f\ i{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{lemma}\ f{\isacharunderscore}incr{\isacharunderscore}lem{\isacharcolon}\ {\isachardoublequote}{\isasymforall}i{\isachardot}\ k\ {\isacharequal}\ f\ i\ {\isasymlongrightarrow}\ i\ {\isasymle}\ f\ i{\isachardoublequote}\isamarkupfalse% % \begin{isamarkuptxt}% \noindent @@ -177,8 +219,8 @@ the same general induction method as for recursion induction (see \S\ref{sec:recdef-induction}):% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ k\ rule{\isacharcolon}\ nat{\isacharunderscore}less{\isacharunderscore}induct{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ k\ rule{\isacharcolon}\ nat{\isacharunderscore}less{\isacharunderscore}induct{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -190,12 +232,12 @@ distinction on \isa{i}. The case \isa{i\ {\isacharequal}\ {\isadigit{0}}} is trivial and we focus on the other case:% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}rule\ allI{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}{\isacharparenleft}case{\isacharunderscore}tac\ i{\isacharparenright}\isanewline \ \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -203,8 +245,15 @@ \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isasymlbrakk}{\isasymforall}m{\isacharless}n{\isachardot}\ {\isasymforall}i{\isachardot}\ m\ {\isacharequal}\ f\ i\ {\isasymlongrightarrow}\ i\ {\isasymle}\ f\ i{\isacharsemicolon}\ i\ {\isacharequal}\ Suc\ nat{\isasymrbrakk}\ {\isasymLongrightarrow}\ n\ {\isacharequal}\ f\ i\ {\isasymlongrightarrow}\ i\ {\isasymle}\ f\ i% \end{isabelle}% \end{isamarkuptxt}% +\isamarkupfalse% +\isacommand{by}{\isacharparenleft}blast\ intro{\isacharbang}{\isacharcolon}\ f{\isacharunderscore}ax\ Suc{\isacharunderscore}leI\ intro{\isacharcolon}\ le{\isacharunderscore}less{\isacharunderscore}trans{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{by}{\isacharparenleft}blast\ intro{\isacharbang}{\isacharcolon}\ f{\isacharunderscore}ax\ Suc{\isacharunderscore}leI\ intro{\isacharcolon}\ le{\isacharunderscore}less{\isacharunderscore}trans{\isacharparenright}\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -235,17 +284,29 @@ The desired result, \isa{i\ {\isasymle}\ f\ i}, follows from \isa{f{\isacharunderscore}incr{\isacharunderscore}lem}:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemmas}\ f{\isacharunderscore}incr\ {\isacharequal}\ f{\isacharunderscore}incr{\isacharunderscore}lem{\isacharbrackleft}rule{\isacharunderscore}format{\isacharcomma}\ OF\ refl{\isacharbrackright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{lemmas}\ f{\isacharunderscore}incr\ {\isacharequal}\ f{\isacharunderscore}incr{\isacharunderscore}lem{\isacharbrackleft}rule{\isacharunderscore}format{\isacharcomma}\ OF\ refl{\isacharbrackright}\isamarkuptrue% % \begin{isamarkuptext}% \noindent The final \isa{refl} gets rid of the premise \isa{{\isacharquery}k\ {\isacharequal}\ f\ {\isacharquery}i}. We could have included this derivation in the original statement of the lemma:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ f{\isacharunderscore}incr{\isacharbrackleft}rule{\isacharunderscore}format{\isacharcomma}\ OF\ refl{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}{\isasymforall}i{\isachardot}\ k\ {\isacharequal}\ f\ i\ {\isasymlongrightarrow}\ i\ {\isasymle}\ f\ i{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{lemma}\ f{\isacharunderscore}incr{\isacharbrackleft}rule{\isacharunderscore}format{\isacharcomma}\ OF\ refl{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}{\isasymforall}i{\isachardot}\ k\ {\isacharequal}\ f\ i\ {\isasymlongrightarrow}\ i\ {\isasymle}\ f\ i{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% % \begin{isamarkuptext}% \begin{exercise} @@ -292,10 +353,16 @@ available for \isa{nat} and want to derive complete induction. We must generalize the statement as shown:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ induct{\isacharunderscore}lem{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isasymAnd}n{\isacharcolon}{\isacharcolon}nat{\isachardot}\ {\isasymforall}m{\isacharless}n{\isachardot}\ P\ m\ {\isasymLongrightarrow}\ P\ n{\isacharparenright}\ {\isasymLongrightarrow}\ {\isasymforall}m{\isacharless}n{\isachardot}\ P\ m{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ n{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ n{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -303,10 +370,17 @@ hypothesis and case \isa{m\ {\isacharequal}\ n} follows from the assumption, again using the induction hypothesis:% \end{isamarkuptxt}% -\ \isamarkuptrue% +\ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}blast{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{by}{\isacharparenleft}blast\ elim{\isacharcolon}\ less{\isacharunderscore}SucE{\isacharparenright}\isamarkupfalse% +\isacommand{by}{\isacharparenleft}blast\ elim{\isacharcolon}\ less{\isacharunderscore}SucE{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -322,10 +396,23 @@ happens automatically when we add the lemma as a new premise to the desired goal:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{theorem}\ nat{\isacharunderscore}less{\isacharunderscore}induct{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isasymAnd}n{\isacharcolon}{\isacharcolon}nat{\isachardot}\ {\isasymforall}m{\isacharless}n{\isachardot}\ P\ m\ {\isasymLongrightarrow}\ P\ n{\isacharparenright}\ {\isasymLongrightarrow}\ P\ n{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}{\isacharparenleft}insert\ induct{\isacharunderscore}lem{\isacharcomma}\ blast{\isacharparenright}\isamarkupfalse% +\isacommand{by}{\isacharparenleft}insert\ induct{\isacharunderscore}lem{\isacharcomma}\ blast{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% HOL already provides the mother of @@ -334,8 +421,19 @@ a special case of \isa{wf{\isacharunderscore}induct} where \isa{r} is \isa{{\isacharless}} on \isa{nat}. The details can be found in theory \isa{Wellfounded_Recursion}.% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Misc/document/Itrev.tex --- a/doc-src/TutorialI/Misc/document/Itrev.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Misc/document/Itrev.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Itrev}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsection{Induction Heuristics% } @@ -45,12 +58,12 @@ \isa{rev} reqires an extra argument where the result is accumulated gradually, using only~\isa{{\isacharhash}}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ itrev\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\isanewline {\isachardoublequote}itrev\ {\isacharbrackleft}{\isacharbrackright}\ \ \ \ \ ys\ {\isacharequal}\ ys{\isachardoublequote}\isanewline -{\isachardoublequote}itrev\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ ys\ {\isacharequal}\ itrev\ xs\ {\isacharparenleft}x{\isacharhash}ys{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +{\isachardoublequote}itrev\ {\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ ys\ {\isacharequal}\ itrev\ xs\ {\isacharparenleft}x{\isacharhash}ys{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -63,15 +76,21 @@ Naturally, we would like to show that \isa{itrev} does indeed reverse its first argument provided the second one is empty:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}itrev\ xs\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ rev\ xs{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{lemma}\ {\isachardoublequote}itrev\ xs\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ rev\ xs{\isachardoublequote}\isamarkupfalse% % \begin{isamarkuptxt}% \noindent There is no choice as to the induction variable, and we immediately simplify:% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -90,10 +109,21 @@ Of course one cannot do this na\"{\i}vely: \isa{itrev\ xs\ ys\ {\isacharequal}\ rev\ xs} is just not true. The correct generalization is% \end{isamarkuptxt}% -\isamarkuptrue% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkupfalse% -\isacommand{lemma}\ {\isachardoublequote}itrev\ xs\ ys\ {\isacharequal}\ rev\ xs\ {\isacharat}\ ys{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}itrev\ xs\ ys\ {\isacharequal}\ rev\ xs\ {\isacharat}\ ys{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -117,10 +147,28 @@ \isa{a\ {\isacharhash}\ ys} instead of \isa{ys}. Hence we prove the theorem for all \isa{ys} instead of a fixed one:% \end{isamarkuptxt}% -\isamarkuptrue% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkupfalse% -\isacommand{lemma}\ {\isachardoublequote}{\isasymforall}ys{\isachardot}\ itrev\ xs\ ys\ {\isacharequal}\ rev\ xs\ {\isacharat}\ ys{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}{\isasymforall}ys{\isachardot}\ itrev\ xs\ ys\ {\isacharequal}\ rev\ xs\ {\isacharat}\ ys{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -156,8 +204,19 @@ to learn about some advanced techniques for inductive proofs.% \index{induction heuristics|)}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Misc/document/Option2.tex --- a/doc-src/TutorialI/Misc/document/Option2.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Misc/document/Option2.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,17 +1,28 @@ % \begin{isabellebody}% \def\isabellecontext{Option{\isadigit{2}}}% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% \indexbold{*option (type)}\indexbold{*None (constant)}% \indexbold{*Some (constant)} Our final datatype is very simple but still eminently useful:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\ {\isacharprime}a\ option\ {\isacharequal}\ None\ {\isacharbar}\ Some\ {\isacharprime}a\isamarkupfalse% +\isamarkupfalse% +\isacommand{datatype}\ {\isacharprime}a\ option\ {\isacharequal}\ None\ {\isacharbar}\ Some\ {\isacharprime}a\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -25,8 +36,19 @@ but it is often simpler to use \isa{option}. For an application see \S\ref{sec:Trie}.% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Misc/document/Plus.tex --- a/doc-src/TutorialI/Misc/document/Plus.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Misc/document/Plus.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,28 +1,74 @@ % \begin{isabellebody}% \def\isabellecontext{Plus}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Define the following addition function% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ plus\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ nat{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\isanewline {\isachardoublequote}plus\ m\ {\isadigit{0}}\ {\isacharequal}\ m{\isachardoublequote}\isanewline -{\isachardoublequote}plus\ m\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ plus\ {\isacharparenleft}Suc\ m{\isacharparenright}\ n{\isachardoublequote}\isamarkupfalse% +{\isachardoublequote}plus\ m\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ plus\ {\isacharparenleft}Suc\ m{\isacharparenright}\ n{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent and prove% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% -\isamarkupfalse% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkupfalse% -\isacommand{lemma}\ {\isachardoublequote}plus\ m\ n\ {\isacharequal}\ m{\isacharplus}n{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}plus\ m\ n\ {\isacharequal}\ m{\isacharplus}n{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Misc/document/Tree.tex --- a/doc-src/TutorialI/Misc/document/Tree.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Misc/document/Tree.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,39 +1,81 @@ % \begin{isabellebody}% \def\isabellecontext{Tree}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Define the datatype of \rmindex{binary trees}:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\ {\isacharprime}a\ tree\ {\isacharequal}\ Tip\ {\isacharbar}\ Node\ {\isachardoublequote}{\isacharprime}a\ tree{\isachardoublequote}\ {\isacharprime}a\ {\isachardoublequote}{\isacharprime}a\ tree{\isachardoublequote}\isamarkupfalse% \isamarkupfalse% -\isamarkupfalse% +\isacommand{datatype}\ {\isacharprime}a\ tree\ {\isacharequal}\ Tip\ {\isacharbar}\ Node\ {\isachardoublequote}{\isacharprime}a\ tree{\isachardoublequote}\ {\isacharprime}a\ {\isachardoublequote}{\isacharprime}a\ tree{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent Define a function \isa{mirror} that mirrors a binary tree by swapping subtrees recursively. Prove% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ mirror{\isacharunderscore}mirror{\isacharcolon}\ {\isachardoublequote}mirror{\isacharparenleft}mirror\ t{\isacharparenright}\ {\isacharequal}\ t{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{lemma}\ mirror{\isacharunderscore}mirror{\isacharcolon}\ {\isachardoublequote}mirror{\isacharparenleft}mirror\ t{\isacharparenright}\ {\isacharequal}\ t{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% % \begin{isamarkuptext}% \noindent Define a function \isa{flatten} that flattens a tree into a list by traversing it in infix order. Prove% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\ {\isachardoublequote}flatten{\isacharparenleft}mirror\ t{\isacharparenright}\ {\isacharequal}\ rev{\isacharparenleft}flatten\ t{\isacharparenright}{\isachardoublequote}\isamarkupfalse% \isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}flatten{\isacharparenleft}mirror\ t{\isacharparenright}\ {\isacharequal}\ rev{\isacharparenleft}flatten\ t{\isacharparenright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Misc/document/Tree2.tex --- a/doc-src/TutorialI/Misc/document/Tree2.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Misc/document/Tree2.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Tree{\isadigit{2}}}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% \noindent In Exercise~\ref{ex:Tree} we defined a function @@ -10,20 +23,52 @@ quadratic. A linear time version of \isa{flatten} again reqires an extra argument, the accumulator:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{consts}\ flatten{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ tree\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequote}\isamarkupfalse% \isamarkupfalse% +\isacommand{consts}\ flatten{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ tree\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent Define \isa{flatten{\isadigit{2}}} and prove% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% -\isamarkupfalse% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkupfalse% -\isacommand{lemma}\ {\isachardoublequote}flatten{\isadigit{2}}\ t\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ flatten\ t{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}flatten{\isadigit{2}}\ t\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ flatten\ t{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Misc/document/appendix.tex --- a/doc-src/TutorialI/Misc/document/appendix.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Misc/document/appendix.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{appendix}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% \begin{table}[htbp] @@ -33,8 +46,19 @@ \end{center} \end{table}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Misc/document/case_exprs.tex --- a/doc-src/TutorialI/Misc/document/case_exprs.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Misc/document/case_exprs.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{case{\isacharunderscore}exprs}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsubsection{Case Expressions% } @@ -64,10 +77,16 @@ distinction over all constructors of the datatype suffices. This is performed by \methdx{case_tac}. Here is a trivial example:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}case\ xs\ of\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymRightarrow}\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbar}\ y{\isacharhash}ys\ {\isasymRightarrow}\ xs{\isacharparenright}\ {\isacharequal}\ xs{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}case{\isacharunderscore}tac\ xs{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}case{\isacharunderscore}tac\ xs{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -79,9 +98,15 @@ \end{isabelle} which is solved automatically:% \end{isamarkuptxt}% +\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}auto{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isamarkupfalse% -\isamarkupfalse% % \begin{isamarkuptext}% Note that we do not need to give a lemma a name if we do not intend to refer @@ -102,8 +127,19 @@ \isa{xs} in the goal. \end{warn}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Misc/document/fakenat.tex --- a/doc-src/TutorialI/Misc/document/fakenat.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Misc/document/fakenat.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,16 +1,40 @@ % \begin{isabellebody}% \def\isabellecontext{fakenat}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% \noindent The type \tydx{nat} of natural numbers is predefined to have the constructors \cdx{0} and~\cdx{Suc}. It behaves as if it were declared like this:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\ nat\ {\isacharequal}\ {\isadigit{0}}\ {\isacharbar}\ Suc\ nat\isamarkupfalse% \isamarkupfalse% +\isacommand{datatype}\ nat\ {\isacharequal}\ {\isadigit{0}}\ {\isacharbar}\ Suc\ nat% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Misc/document/natsum.tex --- a/doc-src/TutorialI/Misc/document/natsum.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Misc/document/natsum.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{natsum}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -11,24 +24,37 @@ \end{isabelle} primitive recursion, for example% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ sum\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat\ {\isasymRightarrow}\ nat{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\ {\isachardoublequote}sum\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequote}\isanewline -\ \ \ \ \ \ \ \ {\isachardoublequote}sum\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ Suc\ n\ {\isacharplus}\ sum\ n{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ \ \ \ {\isachardoublequote}sum\ {\isacharparenleft}Suc\ n{\isacharparenright}\ {\isacharequal}\ Suc\ n\ {\isacharplus}\ sum\ n{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent and induction, for example% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}sum\ n\ {\isacharplus}\ sum\ n\ {\isacharequal}\ n{\isacharasterisk}{\isacharparenleft}Suc\ n{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ n{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \newcommand{\mystar}{*% @@ -84,9 +110,21 @@ (a method introduced below, \S\ref{sec:Simplification}) prove simple arithmetic goals automatically:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}{\isasymlbrakk}\ {\isasymnot}\ m\ {\isacharless}\ n{\isacharsemicolon}\ m\ {\isacharless}\ n\ {\isacharplus}\ {\isacharparenleft}{\isadigit{1}}{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ m\ {\isacharequal}\ n{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{lemma}\ {\isachardoublequote}{\isasymlbrakk}\ {\isasymnot}\ m\ {\isacharless}\ n{\isacharsemicolon}\ m\ {\isacharless}\ n\ {\isacharplus}\ {\isacharparenleft}{\isadigit{1}}{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ m\ {\isacharequal}\ n{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -94,9 +132,21 @@ many logical connectives, and all arithmetic operations apart from addition. In consequence, \isa{auto} and \isa{simp} cannot prove this slightly more complex goal:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}m\ {\isasymnoteq}\ {\isacharparenleft}n{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ {\isasymLongrightarrow}\ m\ {\isacharless}\ n\ {\isasymor}\ n\ {\isacharless}\ m{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{lemma}\ {\isachardoublequote}m\ {\isasymnoteq}\ {\isacharparenleft}n{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ {\isasymLongrightarrow}\ m\ {\isacharless}\ n\ {\isasymor}\ n\ {\isacharless}\ m{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% % \begin{isamarkuptext}% \noindent The method \methdx{arith} is more general. It attempts to @@ -107,19 +157,43 @@ \isa{{\isasymle}} and \isa{{\isacharless}}, and the operations \isa{{\isacharplus}}, \isa{{\isacharminus}}, \isa{min} and \isa{max}. For example,% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}min\ i\ {\isacharparenleft}max\ j\ {\isacharparenleft}k{\isacharasterisk}k{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ max\ {\isacharparenleft}min\ {\isacharparenleft}k{\isacharasterisk}k{\isacharparenright}\ i{\isacharparenright}\ {\isacharparenleft}min\ i\ {\isacharparenleft}j{\isacharcolon}{\isacharcolon}nat{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}arith{\isacharparenright}\isamarkupfalse% -\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}arith{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent succeeds because \isa{k\ {\isacharasterisk}\ k} can be treated as atomic. In contrast,% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}n{\isacharasterisk}n\ {\isacharequal}\ n\ {\isasymLongrightarrow}\ n{\isacharequal}{\isadigit{0}}\ {\isasymor}\ n{\isacharequal}{\isadigit{1}}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{lemma}\ {\isachardoublequote}n{\isacharasterisk}n\ {\isacharequal}\ n\ {\isasymLongrightarrow}\ n{\isacharequal}{\isadigit{0}}\ {\isasymor}\ n{\isacharequal}{\isadigit{1}}{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -139,8 +213,19 @@ super-exponential time and space. \end{warn}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Misc/document/pairs.tex --- a/doc-src/TutorialI/Misc/document/pairs.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Misc/document/pairs.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{pairs}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% \label{sec:pairs}\index{pairs and tuples} @@ -32,8 +45,19 @@ \end{itemize} For more information on pairs and records see Chapter~\ref{ch:more-types}.% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Misc/document/prime_def.tex --- a/doc-src/TutorialI/Misc/document/prime_def.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Misc/document/prime_def.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,8 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{prime{\isacharunderscore}def}% -\isamarkupfalse% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% \begin{warn} @@ -21,8 +33,19 @@ \end{isabelle} \end{warn}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Misc/document/simp.tex --- a/doc-src/TutorialI/Misc/document/simp.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Misc/document/simp.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{simp}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsubsection{Simplification Rules% } @@ -111,12 +124,25 @@ By default, assumptions are part of the simplification process: they are used as simplification rules and are simplified themselves. For example:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isasymlbrakk}\ xs\ {\isacharat}\ zs\ {\isacharequal}\ ys\ {\isacharat}\ xs{\isacharsemicolon}\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharat}\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharat}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ ys\ {\isacharequal}\ zs{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}\ simp\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -126,8 +152,14 @@ In some cases, using the assumptions can lead to nontermination:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}{\isasymforall}x{\isachardot}\ f\ x\ {\isacharequal}\ g\ {\isacharparenleft}f\ {\isacharparenleft}g\ x{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\ f\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ f\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharat}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{lemma}\ {\isachardoublequote}{\isasymforall}x{\isachardot}\ f\ x\ {\isacharequal}\ g\ {\isacharparenleft}f\ {\isacharparenleft}g\ x{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\ f\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ f\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharat}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}\isamarkupfalse% % \begin{isamarkuptxt}% \noindent @@ -137,10 +169,17 @@ nontermination but not this one.) The problem can be circumvented by telling the simplifier to ignore the assumptions:% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}simp\ {\isacharparenleft}no{\isacharunderscore}asm{\isacharparenright}{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -188,24 +227,30 @@ For example, given% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{constdefs}\ xor\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}bool\ {\isasymRightarrow}\ bool\ {\isasymRightarrow}\ bool{\isachardoublequote}\isanewline -\ \ \ \ \ \ \ \ \ {\isachardoublequote}xor\ A\ B\ {\isasymequiv}\ {\isacharparenleft}A\ {\isasymand}\ {\isasymnot}B{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymnot}A\ {\isasymand}\ B{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ \ \ \ \ {\isachardoublequote}xor\ A\ B\ {\isasymequiv}\ {\isacharparenleft}A\ {\isasymand}\ {\isasymnot}B{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymnot}A\ {\isasymand}\ B{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent we may want to prove% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}xor\ A\ {\isacharparenleft}{\isasymnot}A{\isacharparenright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{lemma}\ {\isachardoublequote}xor\ A\ {\isacharparenleft}{\isasymnot}A{\isacharparenright}{\isachardoublequote}\isamarkupfalse% % \begin{isamarkuptxt}% \noindent Typically, we begin by unfolding some definitions: \indexbold{definitions!unfolding}% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}{\isacharparenleft}simp\ only{\isacharcolon}\ xor{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}simp\ only{\isacharcolon}\ xor{\isacharunderscore}def{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -215,11 +260,28 @@ \end{isabelle} can be proved by simplification. Thus we could have proved the lemma outright by% \end{isamarkuptxt}% -\isamarkuptrue% -\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp\ add{\isacharcolon}\ xor{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% -\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}simp\ add{\isacharcolon}\ xor{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -252,20 +314,33 @@ the predefined constant \isa{Let}, expanding \isa{let}-constructs means rewriting with \tdx{Let_def}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}let\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ in\ xs{\isacharat}ys{\isacharat}xs{\isacharparenright}\ {\isacharequal}\ ys{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}simp\ add{\isacharcolon}\ Let{\isacharunderscore}def{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% If, in a particular context, there is no danger of a combinatorial explosion of nested \isa{let}s, you could even simplify with \isa{Let{\isacharunderscore}def} by default:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{declare}\ Let{\isacharunderscore}def\ {\isacharbrackleft}simp{\isacharbrackright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{declare}\ Let{\isacharunderscore}def\ {\isacharbrackleft}simp{\isacharbrackright}\isamarkuptrue% % \isamarkupsubsection{Conditional Simplification Rules% } @@ -276,12 +351,25 @@ So far all examples of rewrite rules were equations. The simplifier also accepts \emph{conditional} equations, for example% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ hd{\isacharunderscore}Cons{\isacharunderscore}tl{\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ \ {\isasymLongrightarrow}\ \ hd\ xs\ {\isacharhash}\ tl\ xs\ {\isacharequal}\ xs{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}case{\isacharunderscore}tac\ xs{\isacharcomma}\ simp{\isacharcomma}\ simp{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -291,9 +379,21 @@ is present as well, the lemma below is proved by plain simplification:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymLongrightarrow}\ hd{\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharhash}\ tl{\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ rev\ xs{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{lemma}\ {\isachardoublequote}xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymLongrightarrow}\ hd{\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharhash}\ tl{\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ rev\ xs{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -315,15 +415,21 @@ are usually proved by case distinction on the boolean condition. Here is an example:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}{\isasymforall}xs{\isachardot}\ if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ rev\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ else\ rev\ xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{lemma}\ {\isachardoublequote}{\isasymforall}xs{\isachardot}\ if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ rev\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ else\ rev\ xs\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}\isamarkupfalse% % \begin{isamarkuptxt}% \noindent The goal can be split by a special method, \methdx{split}:% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}{\isacharparenleft}split\ split{\isacharunderscore}if{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}split\ split{\isacharunderscore}if{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -339,11 +445,23 @@ This splitting idea generalizes from \isa{if} to \sdx{case}. Let us simplify a case analysis over lists:\index{*list.split (theorem)}% \end{isamarkuptxt}% -\isamarkuptrue% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}case\ xs\ of\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymRightarrow}\ zs\ {\isacharbar}\ y{\isacharhash}ys\ {\isasymRightarrow}\ y{\isacharhash}{\isacharparenleft}ys{\isacharat}zs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ xs{\isacharat}zs{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}split\ list{\isachardot}split{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}split\ list{\isachardot}split{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -358,11 +476,28 @@ for adding splitting rules explicitly. The lemma above can be proved in one step by% \end{isamarkuptxt}% -\isamarkuptrue% -\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp\ split{\isacharcolon}\ list{\isachardot}split{\isacharparenright}\isamarkupfalse% -\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}simp\ split{\isacharcolon}\ list{\isachardot}split{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -372,25 +507,36 @@ $t$\isa{{\isachardot}split} which can be declared to be a \bfindex{split rule} either locally as above, or by giving it the \attrdx{split} attribute globally:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{declare}\ list{\isachardot}split\ {\isacharbrackleft}split{\isacharbrackright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{declare}\ list{\isachardot}split\ {\isacharbrackleft}split{\isacharbrackright}\isamarkuptrue% % \begin{isamarkuptext}% \noindent The \isa{split} attribute can be removed with the \isa{del} modifier, either locally% \end{isamarkuptext}% -\isamarkuptrue% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp\ split\ del{\isacharcolon}\ split{\isacharunderscore}if{\isacharparenright}\isamarkupfalse% -\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}simp\ split\ del{\isacharcolon}\ split{\isacharunderscore}if{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent or globally:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{declare}\ list{\isachardot}split\ {\isacharbrackleft}split\ del{\isacharbrackright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{declare}\ list{\isachardot}split\ {\isacharbrackleft}split\ del{\isacharbrackright}\isamarkuptrue% % \begin{isamarkuptext}% Polished proofs typically perform splitting within \isa{simp} rather than @@ -404,10 +550,16 @@ in the assumptions, you have to apply \tdx{split_if_asm} or $t$\isa{{\isachardot}split{\isacharunderscore}asm}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}if\ xs\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ then\ ys\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}\ else\ ys\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}\ {\isasymLongrightarrow}\ xs\ {\isacharat}\ ys\ {\isasymnoteq}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}split\ split{\isacharunderscore}if{\isacharunderscore}asm{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}split\ split{\isacharunderscore}if{\isacharunderscore}asm{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -431,8 +583,14 @@ cases or it is split. \end{warn}% \end{isamarkuptxt}% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isamarkupfalse% % \isamarkupsubsection{Tracing% } @@ -443,11 +601,23 @@ Using the simplifier effectively may take a bit of experimentation. Set the Proof General flag \pgmenu{Isabelle} $>$ \pgmenu{Settings} $>$ \pgmenu{Trace Simplifier} to get a better idea of what is going on:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}rev\ {\isacharbrackleft}a{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isamarkupfalse% -\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}simp{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -583,8 +753,19 @@ to type in lengthy expressions again and again. \end{pgnote}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Misc/document/types.tex --- a/doc-src/TutorialI/Misc/document/types.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Misc/document/types.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,10 +1,23 @@ % \begin{isabellebody}% \def\isabellecontext{types}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \isamarkupfalse% \isacommand{types}\ number\ \ \ \ \ \ \ {\isacharequal}\ nat\isanewline \ \ \ \ \ \ gate\ \ \ \ \ \ \ \ \ {\isacharequal}\ {\isachardoublequote}bool\ {\isasymRightarrow}\ bool\ {\isasymRightarrow}\ bool{\isachardoublequote}\isanewline -\ \ \ \ \ \ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}b{\isacharparenright}alist\ {\isacharequal}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b{\isacharparenright}list{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ \ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}b{\isacharparenright}alist\ {\isacharequal}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b{\isacharparenright}list{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -13,9 +26,9 @@ readability of theories. Synonyms can be used just like any other type. Here, we declare two constants of type \isa{gate}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ nand\ {\isacharcolon}{\isacharcolon}\ gate\isanewline -\ \ \ \ \ \ \ xor\ \ {\isacharcolon}{\isacharcolon}\ gate\isamarkupfalse% +\ \ \ \ \ \ \ xor\ \ {\isacharcolon}{\isacharcolon}\ gate\isamarkuptrue% % \isamarkupsubsection{Constant Definitions% } @@ -26,9 +39,9 @@ The constants \isa{nand} and \isa{xor} above are non-recursive and can be defined directly:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{defs}\ nand{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequote}nand\ A\ B\ {\isasymequiv}\ {\isasymnot}{\isacharparenleft}A\ {\isasymand}\ B{\isacharparenright}{\isachardoublequote}\isanewline -\ \ \ \ \ xor{\isacharunderscore}def{\isacharcolon}\ \ {\isachardoublequote}xor\ A\ B\ \ {\isasymequiv}\ A\ {\isasymand}\ {\isasymnot}B\ {\isasymor}\ {\isasymnot}A\ {\isasymand}\ B{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ xor{\isacharunderscore}def{\isacharcolon}\ \ {\isachardoublequote}xor\ A\ B\ \ {\isasymequiv}\ A\ {\isasymand}\ {\isasymnot}B\ {\isasymor}\ {\isasymnot}A\ {\isasymand}\ B{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent% @@ -45,19 +58,30 @@ \isacommand{defs}. For instance, we can introduce \isa{nand} and \isa{xor} by a single command:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{constdefs}\ nor\ {\isacharcolon}{\isacharcolon}\ gate\isanewline \ \ \ \ \ \ \ \ \ {\isachardoublequote}nor\ A\ B\ {\isasymequiv}\ {\isasymnot}{\isacharparenleft}A\ {\isasymor}\ B{\isacharparenright}{\isachardoublequote}\isanewline \ \ \ \ \ \ \ \ \ \ xor{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ gate\isanewline -\ \ \ \ \ \ \ \ \ {\isachardoublequote}xor{\isadigit{2}}\ A\ B\ {\isasymequiv}\ {\isacharparenleft}A\ {\isasymor}\ B{\isacharparenright}\ {\isasymand}\ {\isacharparenleft}{\isasymnot}A\ {\isasymor}\ {\isasymnot}B{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ \ \ \ \ {\isachardoublequote}xor{\isadigit{2}}\ A\ B\ {\isasymequiv}\ {\isacharparenleft}A\ {\isasymor}\ B{\isacharparenright}\ {\isasymand}\ {\isacharparenleft}{\isasymnot}A\ {\isasymor}\ {\isasymnot}B{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent The default name of each definition is $f$\isa{{\isacharunderscore}def}, where $f$ is the name of the defined constant.% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Recdef/document/Induction.tex --- a/doc-src/TutorialI/Recdef/document/Induction.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Recdef/document/Induction.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Induction}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% Assuming we have defined our function such that Isabelle could prove @@ -19,8 +32,14 @@ for all recursive calls on the right-hand side. Here is a simple example involving the predefined \isa{map} functional on lists:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}map\ f\ {\isacharparenleft}sep{\isacharparenleft}x{\isacharcomma}xs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep{\isacharparenleft}f\ x{\isacharcomma}\ map\ f\ xs{\isacharparenright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{lemma}\ {\isachardoublequote}map\ f\ {\isacharparenleft}sep{\isacharparenleft}x{\isacharcomma}xs{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ sep{\isacharparenleft}f\ x{\isacharcomma}\ map\ f\ xs{\isacharparenright}{\isachardoublequote}\isamarkupfalse% % \begin{isamarkuptxt}% \noindent @@ -28,8 +47,8 @@ is the result of applying \isa{f} to all elements of \isa{xs}. We prove this lemma by recursion induction over \isa{sep}:% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ x\ xs\ rule{\isacharcolon}\ sep{\isachardot}induct{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ x\ xs\ rule{\isacharcolon}\ sep{\isachardot}induct{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -44,10 +63,17 @@ \end{isabelle} The rest is pure simplification:% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}\ simp{\isacharunderscore}all\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% Try proving the above lemma by structural induction, and you find that you @@ -75,8 +101,19 @@ The final case has an induction hypothesis: you may assume that \isa{P} holds for the tail of that list.% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Recdef/document/Nested0.tex --- a/doc-src/TutorialI/Recdef/document/Nested0.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Recdef/document/Nested0.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,14 +1,27 @@ % \begin{isabellebody}% \def\isabellecontext{Nested{\isadigit{0}}}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% \index{datatypes!nested}% In \S\ref{sec:nested-datatype} we defined the datatype of terms% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}b{\isacharparenright}{\isachardoublequote}term{\isachardoublequote}\ {\isacharequal}\ Var\ {\isacharprime}a\ {\isacharbar}\ App\ {\isacharprime}b\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}b{\isacharparenright}term\ list{\isachardoublequote}\isamarkupfalse% +\isamarkupfalse% +\isacommand{datatype}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}b{\isacharparenright}{\isachardoublequote}term{\isachardoublequote}\ {\isacharequal}\ Var\ {\isacharprime}a\ {\isacharbar}\ App\ {\isacharprime}b\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}b{\isacharparenright}term\ list{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -21,9 +34,20 @@ definitions and proofs about nested recursive datatypes. As an example we choose exercise~\ref{ex:trev-trev}:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{consts}\ trev\ \ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}b{\isacharparenright}term\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}b{\isacharparenright}term{\isachardoublequote}\isamarkupfalse% \isamarkupfalse% +\isacommand{consts}\ trev\ \ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}b{\isacharparenright}term\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}b{\isacharparenright}term{\isachardoublequote}% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Recdef/document/Nested1.tex --- a/doc-src/TutorialI/Recdef/document/Nested1.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Recdef/document/Nested1.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Nested{\isadigit{1}}}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -14,10 +27,10 @@ simplifies matters because we are now free to use the recursion equation suggested at the end of \S\ref{sec:nested-datatype}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{recdef}\ trev\ {\isachardoublequote}measure\ size{\isachardoublequote}\isanewline \ {\isachardoublequote}trev\ {\isacharparenleft}Var\ x{\isacharparenright}\ \ \ \ {\isacharequal}\ Var\ x{\isachardoublequote}\isanewline -\ {\isachardoublequote}trev\ {\isacharparenleft}App\ f\ ts{\isacharparenright}\ {\isacharequal}\ App\ f\ {\isacharparenleft}rev{\isacharparenleft}map\ trev\ ts{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ {\isachardoublequote}trev\ {\isacharparenleft}App\ f\ ts{\isacharparenright}\ {\isacharequal}\ App\ f\ {\isacharparenleft}rev{\isacharparenleft}map\ trev\ ts{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -42,8 +55,19 @@ The termination condition is easily proved by induction:% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Recdef/document/Nested2.tex --- a/doc-src/TutorialI/Recdef/document/Nested2.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Recdef/document/Nested2.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,12 +1,37 @@ % \begin{isabellebody}% \def\isabellecontext{Nested{\isadigit{2}}}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory \isanewline +% +\endisadelimtheory \isamarkupfalse% \isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}t\ {\isasymin}\ set\ ts\ {\isasymlongrightarrow}\ size\ t\ {\isacharless}\ Suc{\isacharparenleft}term{\isacharunderscore}list{\isacharunderscore}size\ ts{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}{\isacharparenleft}induct{\isacharunderscore}tac\ ts{\isacharcomma}\ auto{\isacharparenright}\isamarkupfalse% -\isamarkupfalse% +\isacommand{by}{\isacharparenleft}induct{\isacharunderscore}tac\ ts{\isacharcomma}\ auto{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -17,10 +42,16 @@ induction schema for type \isa{term} and can use the simpler one arising from \isa{trev}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}trev{\isacharparenleft}trev\ t{\isacharparenright}\ {\isacharequal}\ t{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ t\ rule{\isacharcolon}\ trev{\isachardot}induct{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ t\ rule{\isacharcolon}\ trev{\isachardot}induct{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -31,8 +62,15 @@ \end{isabelle} Both the base case and the induction step fall to simplification:% \end{isamarkuptxt}% +\isamarkupfalse% +\isacommand{by}{\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ rev{\isacharunderscore}map\ sym{\isacharbrackleft}OF\ map{\isacharunderscore}compose{\isacharbrackright}\ cong{\isacharcolon}\ map{\isacharunderscore}cong{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{by}{\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ rev{\isacharunderscore}map\ sym{\isacharbrackleft}OF\ map{\isacharunderscore}compose{\isacharbrackright}\ cong{\isacharcolon}\ map{\isacharunderscore}cong{\isacharparenright}\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -76,17 +114,15 @@ congruence rules, you can append a hint after the end of the recursion equations:\cmmdx{hints}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% -{\isacharparenleft}\isakeyword{hints}\ recdef{\isacharunderscore}cong{\isacharcolon}\ map{\isacharunderscore}cong{\isacharparenright}\isamarkupfalse% +{\isacharparenleft}\isakeyword{hints}\ recdef{\isacharunderscore}cong{\isacharcolon}\ map{\isacharunderscore}cong{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptext}% \noindent Or you can declare them globally by giving them the \attrdx{recdef_cong} attribute:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{declare}\ map{\isacharunderscore}cong{\isacharbrackleft}recdef{\isacharunderscore}cong{\isacharbrackright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{declare}\ map{\isacharunderscore}cong{\isacharbrackleft}recdef{\isacharunderscore}cong{\isacharbrackright}\isamarkuptrue% % \begin{isamarkuptext}% The \isa{cong} and \isa{recdef{\isacharunderscore}cong} attributes are @@ -96,8 +132,19 @@ %For example the weak congruence rules for if and case would prevent %recdef from generating sensible termination conditions.% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Recdef/document/examples.tex --- a/doc-src/TutorialI/Recdef/document/examples.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Recdef/document/examples.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,18 +1,31 @@ % \begin{isabellebody}% \def\isabellecontext{examples}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% Here is a simple example, the \rmindex{Fibonacci function}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ fib\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat\ {\isasymRightarrow}\ nat{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{recdef}\ fib\ {\isachardoublequote}measure{\isacharparenleft}{\isasymlambda}n{\isachardot}\ n{\isacharparenright}{\isachardoublequote}\isanewline \ \ {\isachardoublequote}fib\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{0}}{\isachardoublequote}\isanewline \ \ {\isachardoublequote}fib\ {\isacharparenleft}Suc\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ {\isadigit{1}}{\isachardoublequote}\isanewline -\ \ {\isachardoublequote}fib\ {\isacharparenleft}Suc{\isacharparenleft}Suc\ x{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ fib\ x\ {\isacharplus}\ fib\ {\isacharparenleft}Suc\ x{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}fib\ {\isacharparenleft}Suc{\isacharparenleft}Suc\ x{\isacharparenright}{\isacharparenright}\ {\isacharequal}\ fib\ x\ {\isacharplus}\ fib\ {\isacharparenleft}Suc\ x{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -29,13 +42,13 @@ Slightly more interesting is the insertion of a fixed element between any two elements of a list:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ sep\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{recdef}\ sep\ {\isachardoublequote}measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}a{\isacharcomma}xs{\isacharparenright}{\isachardot}\ length\ xs{\isacharparenright}{\isachardoublequote}\isanewline \ \ {\isachardoublequote}sep{\isacharparenleft}a{\isacharcomma}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ \ \ \ \ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}\isanewline \ \ {\isachardoublequote}sep{\isacharparenleft}a{\isacharcomma}\ {\isacharbrackleft}x{\isacharbrackright}{\isacharparenright}\ \ \ \ {\isacharequal}\ {\isacharbrackleft}x{\isacharbrackright}{\isachardoublequote}\isanewline -\ \ {\isachardoublequote}sep{\isacharparenleft}a{\isacharcomma}\ x{\isacharhash}y{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ x\ {\isacharhash}\ a\ {\isacharhash}\ sep{\isacharparenleft}a{\isacharcomma}y{\isacharhash}zs{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}sep{\isacharparenleft}a{\isacharcomma}\ x{\isacharhash}y{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ x\ {\isacharhash}\ a\ {\isacharhash}\ sep{\isacharparenleft}a{\isacharcomma}y{\isacharhash}zs{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -47,23 +60,23 @@ Pattern matching\index{pattern matching!and \isacommand{recdef}} need not be exhaustive:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ last\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{recdef}\ last\ {\isachardoublequote}measure\ {\isacharparenleft}{\isasymlambda}xs{\isachardot}\ length\ xs{\isacharparenright}{\isachardoublequote}\isanewline \ \ {\isachardoublequote}last\ {\isacharbrackleft}x{\isacharbrackright}\ \ \ \ \ \ {\isacharequal}\ x{\isachardoublequote}\isanewline -\ \ {\isachardoublequote}last\ {\isacharparenleft}x{\isacharhash}y{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ last\ {\isacharparenleft}y{\isacharhash}zs{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}last\ {\isacharparenleft}x{\isacharhash}y{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ last\ {\isacharparenleft}y{\isacharhash}zs{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% Overlapping patterns are disambiguated by taking the order of equations into account, just as in functional programming:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ sep{\isadigit{1}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{recdef}\ sep{\isadigit{1}}\ {\isachardoublequote}measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}a{\isacharcomma}xs{\isacharparenright}{\isachardot}\ length\ xs{\isacharparenright}{\isachardoublequote}\isanewline \ \ {\isachardoublequote}sep{\isadigit{1}}{\isacharparenleft}a{\isacharcomma}\ x{\isacharhash}y{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ x\ {\isacharhash}\ a\ {\isacharhash}\ sep{\isadigit{1}}{\isacharparenleft}a{\isacharcomma}y{\isacharhash}zs{\isacharparenright}{\isachardoublequote}\isanewline -\ \ {\isachardoublequote}sep{\isadigit{1}}{\isacharparenleft}a{\isacharcomma}\ xs{\isacharparenright}\ \ \ \ \ {\isacharequal}\ xs{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}sep{\isadigit{1}}{\isacharparenleft}a{\isacharcomma}\ xs{\isacharparenright}\ \ \ \ \ {\isacharequal}\ xs{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -82,26 +95,37 @@ arguments as in the following definition: \end{warn}% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ sep{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{recdef}\ sep{\isadigit{2}}\ {\isachardoublequote}measure\ length{\isachardoublequote}\isanewline \ \ {\isachardoublequote}sep{\isadigit{2}}\ {\isacharparenleft}x{\isacharhash}y{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}{\isasymlambda}a{\isachardot}\ x\ {\isacharhash}\ a\ {\isacharhash}\ sep{\isadigit{2}}\ {\isacharparenleft}y{\isacharhash}zs{\isacharparenright}\ a{\isacharparenright}{\isachardoublequote}\isanewline -\ \ {\isachardoublequote}sep{\isadigit{2}}\ xs\ \ \ \ \ \ \ {\isacharequal}\ {\isacharparenleft}{\isasymlambda}a{\isachardot}\ xs{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}sep{\isadigit{2}}\ xs\ \ \ \ \ \ \ {\isacharequal}\ {\isacharparenleft}{\isasymlambda}a{\isachardot}\ xs{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% Because of its pattern-matching syntax, \isacommand{recdef} is also useful for the definition of non-recursive functions, where the termination measure degenerates to the empty set \isa{{\isacharbraceleft}{\isacharbraceright}}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ swap{\isadigit{1}}{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{recdef}\ swap{\isadigit{1}}{\isadigit{2}}\ {\isachardoublequote}{\isacharbraceleft}{\isacharbraceright}{\isachardoublequote}\isanewline \ \ {\isachardoublequote}swap{\isadigit{1}}{\isadigit{2}}\ {\isacharparenleft}x{\isacharhash}y{\isacharhash}zs{\isacharparenright}\ {\isacharequal}\ y{\isacharhash}x{\isacharhash}zs{\isachardoublequote}\isanewline \ \ {\isachardoublequote}swap{\isadigit{1}}{\isadigit{2}}\ zs\ \ \ \ \ \ \ {\isacharequal}\ zs{\isachardoublequote}\isanewline -\isamarkupfalse% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Recdef/document/simplification.tex --- a/doc-src/TutorialI/Recdef/document/simplification.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Recdef/document/simplification.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{simplification}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% Once we have proved all the termination conditions, the \isacommand{recdef} @@ -12,11 +25,11 @@ \index{*if expressions!splitting of} Let us look at an example:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ gcd\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat{\isasymtimes}nat\ {\isasymRightarrow}\ nat{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{recdef}\ gcd\ {\isachardoublequote}measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}m{\isacharcomma}n{\isacharparenright}{\isachardot}n{\isacharparenright}{\isachardoublequote}\isanewline -\ \ {\isachardoublequote}gcd\ {\isacharparenleft}m{\isacharcomma}\ n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ n{\isacharequal}{\isadigit{0}}\ then\ m\ else\ gcd{\isacharparenleft}n{\isacharcomma}\ m\ mod\ n{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}gcd\ {\isacharparenleft}m{\isacharcomma}\ n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}if\ n{\isacharequal}{\isadigit{0}}\ then\ m\ else\ gcd{\isacharparenleft}n{\isacharcomma}\ m\ mod\ n{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -60,12 +73,12 @@ rather than \isa{if} on the right. In the case of \isa{simplification{\isachardot}gcd} the following alternative definition suggests itself:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ gcd{\isadigit{1}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat{\isasymtimes}nat\ {\isasymRightarrow}\ nat{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{recdef}\ gcd{\isadigit{1}}\ {\isachardoublequote}measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}m{\isacharcomma}n{\isacharparenright}{\isachardot}n{\isacharparenright}{\isachardoublequote}\isanewline \ \ {\isachardoublequote}gcd{\isadigit{1}}\ {\isacharparenleft}m{\isacharcomma}\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ m{\isachardoublequote}\isanewline -\ \ {\isachardoublequote}gcd{\isadigit{1}}\ {\isacharparenleft}m{\isacharcomma}\ n{\isacharparenright}\ {\isacharequal}\ gcd{\isadigit{1}}{\isacharparenleft}n{\isacharcomma}\ m\ mod\ n{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}gcd{\isadigit{1}}\ {\isacharparenleft}m{\isacharcomma}\ n{\isacharparenright}\ {\isacharequal}\ gcd{\isadigit{1}}{\isacharparenleft}n{\isacharcomma}\ m\ mod\ n{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -76,11 +89,11 @@ A simple alternative is to replace \isa{if} by \isa{case}, which is also available for \isa{bool} and is not split automatically:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ gcd{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat{\isasymtimes}nat\ {\isasymRightarrow}\ nat{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{recdef}\ gcd{\isadigit{2}}\ {\isachardoublequote}measure\ {\isacharparenleft}{\isasymlambda}{\isacharparenleft}m{\isacharcomma}n{\isacharparenright}{\isachardot}n{\isacharparenright}{\isachardoublequote}\isanewline -\ \ {\isachardoublequote}gcd{\isadigit{2}}{\isacharparenleft}m{\isacharcomma}n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}case\ n{\isacharequal}{\isadigit{0}}\ of\ True\ {\isasymRightarrow}\ m\ {\isacharbar}\ False\ {\isasymRightarrow}\ gcd{\isadigit{2}}{\isacharparenleft}n{\isacharcomma}m\ mod\ n{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}gcd{\isadigit{2}}{\isacharparenleft}m{\isacharcomma}n{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}case\ n{\isacharequal}{\isadigit{0}}\ of\ True\ {\isasymRightarrow}\ m\ {\isacharbar}\ False\ {\isasymRightarrow}\ gcd{\isadigit{2}}{\isacharparenleft}n{\isacharcomma}m\ mod\ n{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -91,29 +104,66 @@ derived conditional ones. For \isa{simplification{\isachardot}gcd} it means we have to prove these lemmas:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}gcd\ {\isacharparenleft}m{\isacharcomma}\ {\isadigit{0}}{\isacharparenright}\ {\isacharequal}\ m{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isanewline +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof \isanewline \isamarkupfalse% \isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}n\ {\isasymnoteq}\ {\isadigit{0}}\ {\isasymLongrightarrow}\ gcd{\isacharparenleft}m{\isacharcomma}\ n{\isacharparenright}\ {\isacharequal}\ gcd{\isacharparenleft}n{\isacharcomma}\ m\ mod\ n{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Simplification terminates for these proofs because the condition of the \isa{if} simplifies to \isa{True} or \isa{False}. Now we can disable the original simplification rule:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{declare}\ gcd{\isachardot}simps\ {\isacharbrackleft}simp\ del{\isacharbrackright}\isanewline -\isamarkupfalse% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Recdef/document/termination.tex --- a/doc-src/TutorialI/Recdef/document/termination.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Recdef/document/termination.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{termination}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% When a function~$f$ is defined via \isacommand{recdef}, Isabelle tries to prove @@ -17,12 +30,12 @@ Isabelle may fail to prove the termination condition for some recursive call. Let us try to define Quicksort:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ qs\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat\ list\ {\isasymRightarrow}\ nat\ list{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{recdef}\ qs\ {\isachardoublequote}measure\ length{\isachardoublequote}\isanewline \ {\isachardoublequote}qs\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}\isanewline -\ {\isachardoublequote}qs{\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ qs{\isacharparenleft}filter\ {\isacharparenleft}{\isasymlambda}y{\isachardot}\ y{\isasymle}x{\isacharparenright}\ xs{\isacharparenright}\ {\isacharat}\ {\isacharbrackleft}x{\isacharbrackright}\ {\isacharat}\ qs{\isacharparenleft}filter\ {\isacharparenleft}{\isasymlambda}y{\isachardot}\ x{\isacharless}y{\isacharparenright}\ xs{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ {\isachardoublequote}qs{\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ qs{\isacharparenleft}filter\ {\isacharparenleft}{\isasymlambda}y{\isachardot}\ y{\isasymle}x{\isacharparenright}\ xs{\isacharparenright}\ {\isacharat}\ {\isacharbrackleft}x{\isacharbrackright}\ {\isacharat}\ qs{\isacharparenleft}filter\ {\isacharparenleft}{\isasymlambda}y{\isachardot}\ x{\isacharless}y{\isacharparenright}\ xs{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent where \isa{filter} is predefined and \isa{filter\ P\ xs} @@ -46,14 +59,11 @@ \attrdx{recdef_simp} attribute says to use \isa{less{\isacharunderscore}Suc{\isacharunderscore}eq{\isacharunderscore}le} as a simplification rule.\cmmdx{hints}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% \isamarkupfalse% \isacommand{recdef}\ qs\ {\isachardoublequote}measure\ length{\isachardoublequote}\isanewline \ {\isachardoublequote}qs\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}\isanewline \ {\isachardoublequote}qs{\isacharparenleft}x{\isacharhash}xs{\isacharparenright}\ {\isacharequal}\ qs{\isacharparenleft}filter\ {\isacharparenleft}{\isasymlambda}y{\isachardot}\ y{\isasymle}x{\isacharparenright}\ xs{\isacharparenright}\ {\isacharat}\ {\isacharbrackleft}x{\isacharbrackright}\ {\isacharat}\ qs{\isacharparenleft}filter\ {\isacharparenleft}{\isasymlambda}y{\isachardot}\ x{\isacharless}y{\isacharparenright}\ xs{\isacharparenright}{\isachardoublequote}\isanewline -{\isacharparenleft}\isakeyword{hints}\ recdef{\isacharunderscore}simp{\isacharcolon}\ less{\isacharunderscore}Suc{\isacharunderscore}eq{\isacharunderscore}le{\isacharparenright}\isamarkupfalse% -\isamarkupfalse% +{\isacharparenleft}\isakeyword{hints}\ recdef{\isacharunderscore}simp{\isacharcolon}\ less{\isacharunderscore}Suc{\isacharunderscore}eq{\isacharunderscore}le{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -62,12 +72,25 @@ simplification rules. Thus we can automatically prove results such as this one:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{theorem}\ {\isachardoublequote}qs{\isacharbrackleft}{\isadigit{2}}{\isacharcomma}{\isadigit{3}}{\isacharcomma}{\isadigit{0}}{\isacharbrackright}\ {\isacharequal}\ qs{\isacharbrackleft}{\isadigit{3}}{\isacharcomma}{\isadigit{0}}{\isacharcomma}{\isadigit{2}}{\isacharbrackright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -78,8 +101,19 @@ \isacommand{hint} is not necessary. But in the case of \isa{less{\isacharunderscore}Suc{\isacharunderscore}eq{\isacharunderscore}le} this would be of dubious value.% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Rules/document/find2.tex --- a/doc-src/TutorialI/Rules/document/find2.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Rules/document/find2.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,8 +1,26 @@ % \begin{isabellebody}% \def\isabellecontext{find{\isadigit{2}}}% -\isamarkupfalse% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkuptrue% % \begin{isamarkuptxt}% \index{finding theorems}\index{searching theorems} In @@ -22,9 +40,20 @@ the very theorem you are trying to prove is already in the database. Given the goal% \end{isamarkuptxt}% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isamarkupfalse% -\isamarkupfalse% % \begin{isamarkuptxt}% \vspace{-\bigskipamount} @@ -45,9 +74,26 @@ \texttt{dest} is analogous to \texttt{intro} but takes the assumptions into account, too.% \end{isamarkuptxt}% -\isamarkuptrue% -\isamarkupfalse% -\isamarkupfalse% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/ToyList/document/ToyList.tex --- a/doc-src/TutorialI/ToyList/document/ToyList.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/ToyList/document/ToyList.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,9 +1,23 @@ % \begin{isabellebody}% \def\isabellecontext{ToyList}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isamarkupfalse% \isacommand{theory}\ ToyList\isanewline \isakeyword{imports}\ PreList\isanewline -\isakeyword{begin}\isamarkupfalse% +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -14,9 +28,9 @@ theory that contains pretty much everything but lists, thus avoiding ambiguities caused by defining lists twice.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{datatype}\ {\isacharprime}a\ list\ {\isacharequal}\ Nil\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}{\isachardoublequote}{\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}{\isacharparenright}\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ Cons\ {\isacharprime}a\ {\isachardoublequote}{\isacharprime}a\ list{\isachardoublequote}\ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}\isakeyword{infixr}\ {\isachardoublequote}{\isacharhash}{\isachardoublequote}\ {\isadigit{6}}{\isadigit{5}}{\isacharparenright}\isamarkupfalse% +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ Cons\ {\isacharprime}a\ {\isachardoublequote}{\isacharprime}a\ list{\isachardoublequote}\ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}\isakeyword{infixr}\ {\isachardoublequote}{\isacharhash}{\isachardoublequote}\ {\isadigit{6}}{\isadigit{5}}{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -49,9 +63,9 @@ \end{warn} Next, two functions \isa{app} and \cdx{rev} are declared:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ app\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequote}\ \ \ {\isacharparenleft}\isakeyword{infixr}\ {\isachardoublequote}{\isacharat}{\isachardoublequote}\ {\isadigit{6}}{\isadigit{5}}{\isacharparenright}\isanewline -\ \ \ \ \ \ \ rev\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ \ \ rev\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -64,7 +78,7 @@ \isa{xs\ {\isacharat}\ ys}\index{$HOL2list@\isa{\at}|bold} becomes the preferred form. Both functions are defined recursively:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{primrec}\isanewline {\isachardoublequote}{\isacharbrackleft}{\isacharbrackright}\ {\isacharat}\ ys\ \ \ \ \ \ \ {\isacharequal}\ ys{\isachardoublequote}\isanewline {\isachardoublequote}{\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ {\isacharat}\ ys\ {\isacharequal}\ x\ {\isacharhash}\ {\isacharparenleft}xs\ {\isacharat}\ ys{\isacharparenright}{\isachardoublequote}\isanewline @@ -72,7 +86,7 @@ \isamarkupfalse% \isacommand{primrec}\isanewline {\isachardoublequote}rev\ {\isacharbrackleft}{\isacharbrackright}\ \ \ \ \ \ \ \ {\isacharequal}\ {\isacharbrackleft}{\isacharbrackright}{\isachardoublequote}\isanewline -{\isachardoublequote}rev\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ \ {\isacharequal}\ {\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharat}\ {\isacharparenleft}x\ {\isacharhash}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +{\isachardoublequote}rev\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ \ {\isacharequal}\ {\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharat}\ {\isacharparenleft}x\ {\isacharhash}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent\index{*rev (constant)|(}\index{append function|(} @@ -112,8 +126,8 @@ To lessen this burden, quotation marks around a single identifier can be dropped, unless the identifier happens to be a keyword, as in% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{consts}\ {\isachardoublequote}end{\isachardoublequote}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\isamarkupfalse% +\isamarkupfalse% +\isacommand{consts}\ {\isachardoublequote}end{\isachardoublequote}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -134,8 +148,14 @@ Our goal is to show that reversing a list twice produces the original list.% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{theorem}\ rev{\isacharunderscore}rev\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}rev{\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ xs{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{theorem}\ rev{\isacharunderscore}rev\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}rev{\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ xs{\isachardoublequote}\isamarkupfalse% % \begin{isamarkuptxt}% \index{theorem@\isacommand {theorem} (command)|bold}% @@ -176,8 +196,8 @@ defined functions are best established by induction. In this case there is nothing obvious except induction on \isa{xs}:% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent\index{*induct_tac (method)}% @@ -212,8 +232,8 @@ Let us try to solve both goals automatically:% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -228,8 +248,14 @@ \end{isabelle} In order to simplify this subgoal further, a lemma suggests itself.% \end{isamarkuptxt}% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isamarkupfalse% % \isamarkupsubsubsection{First Lemma% } @@ -240,8 +266,14 @@ After abandoning the above proof attempt (at the shell level type \commdx{oops}) we start a new proof:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ rev{\isacharunderscore}app\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}rev{\isacharparenleft}xs\ {\isacharat}\ ys{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}rev\ ys{\isacharparenright}\ {\isacharat}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{lemma}\ rev{\isacharunderscore}app\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}rev{\isacharparenleft}xs\ {\isacharat}\ ys{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}rev\ ys{\isacharparenright}\ {\isacharat}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isachardoublequote}\isamarkupfalse% % \begin{isamarkuptxt}% \noindent The keywords \commdx{theorem} and @@ -253,15 +285,15 @@ \isa{ys}. Because \isa{{\isacharat}} is defined by recursion on the first argument, \isa{xs} is the correct one:% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent This time not even the base case is solved automatically:% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -271,8 +303,14 @@ first. In the future the step of abandoning an incomplete proof before embarking on the proof of a lemma usually remains implicit.% \end{isamarkuptxt}% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isamarkupfalse% % \isamarkupsubsubsection{Second Lemma% } @@ -281,12 +319,18 @@ \begin{isamarkuptext}% We again try the canonical proof procedure:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ app{\isacharunderscore}Nil{\isadigit{2}}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}xs\ {\isacharat}\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ xs{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -297,8 +341,15 @@ \end{isabelle} We still need to confirm that the proof is now finished:% \end{isamarkuptxt}% +\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{done}\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -315,12 +366,18 @@ Going back to the proof of the first lemma% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ rev{\isacharunderscore}app\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}rev{\isacharparenleft}xs\ {\isacharat}\ ys{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}rev\ ys{\isacharparenright}\ {\isacharat}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -339,8 +396,14 @@ \end{isabelle} and the missing lemma is associativity of \isa{{\isacharat}}.% \end{isamarkuptxt}% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isamarkupfalse% % \isamarkupsubsubsection{Third Lemma% } @@ -350,40 +413,79 @@ Abandoning the previous attempt, the canonical proof procedure succeeds without further ado.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ app{\isacharunderscore}assoc\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}xs\ {\isacharat}\ ys{\isacharparenright}\ {\isacharat}\ zs\ {\isacharequal}\ xs\ {\isacharat}\ {\isacharparenleft}ys\ {\isacharat}\ zs{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Now we can prove the first lemma:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ rev{\isacharunderscore}app\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}rev{\isacharparenleft}xs\ {\isacharat}\ ys{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}rev\ ys{\isacharparenright}\ {\isacharat}\ {\isacharparenleft}rev\ xs{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Finally, we prove our main theorem:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{theorem}\ rev{\isacharunderscore}rev\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}rev{\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ xs{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ xs{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}{\isacharparenleft}auto{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -391,9 +493,21 @@ we are finished with its development:% \index{*rev (constant)|)}\index{append function|)}% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{end}\isanewline +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory \isamarkupfalse% +\isacommand{end}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Trie/document/Trie.tex --- a/doc-src/TutorialI/Trie/document/Trie.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Trie/document/Trie.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Trie}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% To minimize running time, each node of a trie should contain an array that maps @@ -10,8 +23,8 @@ list of (letter,trie) pairs. Abstracting over the alphabet \isa{{\isacharprime}a} and the values \isa{{\isacharprime}v} we define a trie as follows:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}trie\ {\isacharequal}\ Trie\ \ {\isachardoublequote}{\isacharprime}v\ option{\isachardoublequote}\ \ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a\ {\isacharasterisk}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}trie{\isacharparenright}list{\isachardoublequote}\isamarkupfalse% +\isamarkupfalse% +\isacommand{datatype}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}trie\ {\isacharequal}\ Trie\ \ {\isachardoublequote}{\isacharprime}v\ option{\isachardoublequote}\ \ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a\ {\isacharasterisk}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}trie{\isacharparenright}list{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -21,25 +34,25 @@ which is fine because products are datatypes as well. We define two selector functions:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ value\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}trie\ {\isasymRightarrow}\ {\isacharprime}v\ option{\isachardoublequote}\isanewline \ \ \ \ \ \ \ alist\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}trie\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isacharasterisk}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}trie{\isacharparenright}list{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\ {\isachardoublequote}value{\isacharparenleft}Trie\ ov\ al{\isacharparenright}\ {\isacharequal}\ ov{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{primrec}\ {\isachardoublequote}alist{\isacharparenleft}Trie\ ov\ al{\isacharparenright}\ {\isacharequal}\ al{\isachardoublequote}\isamarkupfalse% +\isacommand{primrec}\ {\isachardoublequote}alist{\isacharparenleft}Trie\ ov\ al{\isacharparenright}\ {\isacharequal}\ al{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent Association lists come with a generic lookup function. Its result involves type \isa{option} because a lookup can fail:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ \ \ assoc\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}key\ {\isacharasterisk}\ {\isacharprime}val{\isacharparenright}list\ {\isasymRightarrow}\ {\isacharprime}key\ {\isasymRightarrow}\ {\isacharprime}val\ option{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\ {\isachardoublequote}assoc\ {\isacharbrackleft}{\isacharbrackright}\ x\ {\isacharequal}\ None{\isachardoublequote}\isanewline \ \ \ \ \ \ \ \ {\isachardoublequote}assoc\ {\isacharparenleft}p{\isacharhash}ps{\isacharparenright}\ x\ {\isacharequal}\isanewline -\ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}let\ {\isacharparenleft}a{\isacharcomma}b{\isacharparenright}\ {\isacharequal}\ p\ in\ if\ a{\isacharequal}x\ then\ Some\ b\ else\ assoc\ ps\ x{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}let\ {\isacharparenleft}a{\isacharcomma}b{\isacharparenright}\ {\isacharequal}\ p\ in\ if\ a{\isacharequal}x\ then\ Some\ b\ else\ assoc\ ps\ x{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% Now we can define the lookup function for tries. It descends into the trie @@ -47,32 +60,45 @@ recursion on lists is simpler than on tries, let us express this as primitive recursion on the search string argument:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ \ \ lookup\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}trie\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}v\ option{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\ {\isachardoublequote}lookup\ t\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ value\ t{\isachardoublequote}\isanewline \ \ \ \ \ \ \ \ {\isachardoublequote}lookup\ t\ {\isacharparenleft}a{\isacharhash}as{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}case\ assoc\ {\isacharparenleft}alist\ t{\isacharparenright}\ a\ of\isanewline \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ None\ {\isasymRightarrow}\ None\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ Some\ at\ {\isasymRightarrow}\ lookup\ at\ as{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ Some\ at\ {\isasymRightarrow}\ lookup\ at\ as{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% As a first simple property we prove that looking up a string in the empty trie \isa{Trie\ None\ {\isacharbrackleft}{\isacharbrackright}} always returns \isa{None}. The proof merely distinguishes the two cases whether the search string is empty or not:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}lookup\ {\isacharparenleft}Trie\ None\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ as\ {\isacharequal}\ None{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}case{\isacharunderscore}tac\ as{\isacharcomma}\ simp{\isacharunderscore}all{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% Things begin to get interesting with the definition of an update function that adds a new (string, value) pair to a trie, overwriting the old value associated with that string:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ update\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}trie\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}v\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}trie{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\isanewline @@ -80,7 +106,7 @@ \ \ {\isachardoublequote}update\ t\ {\isacharparenleft}a{\isacharhash}as{\isacharparenright}\ v\ {\isacharequal}\isanewline \ \ \ \ \ {\isacharparenleft}let\ tt\ {\isacharequal}\ {\isacharparenleft}case\ assoc\ {\isacharparenleft}alist\ t{\isacharparenright}\ a\ of\isanewline \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ None\ {\isasymRightarrow}\ Trie\ None\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbar}\ Some\ at\ {\isasymRightarrow}\ at{\isacharparenright}\isanewline -\ \ \ \ \ \ in\ Trie\ {\isacharparenleft}value\ t{\isacharparenright}\ {\isacharparenleft}{\isacharparenleft}a{\isacharcomma}update\ tt\ as\ v{\isacharparenright}\ {\isacharhash}\ alist\ t{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ \ in\ Trie\ {\isacharparenleft}value\ t{\isacharparenright}\ {\isacharparenleft}{\isacharparenleft}a{\isacharcomma}update\ tt\ as\ v{\isacharparenright}\ {\isacharhash}\ alist\ t{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -95,8 +121,8 @@ expand all \isa{let}s and to split all \isa{case}-constructs over options:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{declare}\ Let{\isacharunderscore}def{\isacharbrackleft}simp{\isacharbrackright}\ option{\isachardot}split{\isacharbrackleft}split{\isacharbrackright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{declare}\ Let{\isacharunderscore}def{\isacharbrackleft}simp{\isacharbrackright}\ option{\isachardot}split{\isacharbrackleft}split{\isacharbrackright}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -107,9 +133,15 @@ Our main goal is to prove the correct interaction of \isa{update} and \isa{lookup}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{theorem}\ {\isachardoublequote}{\isasymforall}t\ v\ bs{\isachardot}\ lookup\ {\isacharparenleft}update\ t\ as\ v{\isacharparenright}\ bs\ {\isacharequal}\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}if\ as{\isacharequal}bs\ then\ Some\ v\ else\ lookup\ t\ bs{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}if\ as{\isacharequal}bs\ then\ Some\ v\ else\ lookup\ t\ bs{\isacharparenright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -121,8 +153,8 @@ \isa{as} is instantiated. The start of the proof is conventional:% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ as{\isacharcomma}\ auto{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ as{\isacharcomma}\ auto{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -136,10 +168,17 @@ well now. It turns out that instead of induction, case distinction suffices:% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}case{\isacharunderscore}tac{\isacharbrackleft}{\isacharbang}{\isacharbrackright}\ bs{\isacharcomma}\ auto{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -178,42 +217,84 @@ with \isa{{\isacharprime}a\ {\isasymrightharpoonup}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}\ {\isacharprime}v{\isacharparenright}\ trie}. \end{exercise}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% -\isamarkupfalse% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Types/document/Axioms.tex --- a/doc-src/TutorialI/Types/document/Axioms.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Types/document/Axioms.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Axioms}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsubsection{Axioms% } @@ -23,12 +36,12 @@ A \emph{partial order} is a subclass of \isa{ordrel} where certain axioms need to hold:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{axclass}\ parord\ {\isacharless}\ ordrel\isanewline refl{\isacharcolon}\ \ \ \ {\isachardoublequote}x\ {\isacharless}{\isacharless}{\isacharequal}\ x{\isachardoublequote}\isanewline trans{\isacharcolon}\ \ \ {\isachardoublequote}{\isasymlbrakk}\ x\ {\isacharless}{\isacharless}{\isacharequal}\ y{\isacharsemicolon}\ y\ {\isacharless}{\isacharless}{\isacharequal}\ z\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ x\ {\isacharless}{\isacharless}{\isacharequal}\ z{\isachardoublequote}\isanewline antisym{\isacharcolon}\ {\isachardoublequote}{\isasymlbrakk}\ x\ {\isacharless}{\isacharless}{\isacharequal}\ y{\isacharsemicolon}\ y\ {\isacharless}{\isacharless}{\isacharequal}\ x\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ x\ {\isacharequal}\ y{\isachardoublequote}\isanewline -less{\isacharunderscore}le{\isacharcolon}\ {\isachardoublequote}x\ {\isacharless}{\isacharless}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}{\isacharless}{\isacharequal}\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +less{\isacharunderscore}le{\isacharcolon}\ {\isachardoublequote}x\ {\isacharless}{\isacharless}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}{\isacharless}{\isacharequal}\ y\ {\isasymand}\ x\ {\isasymnoteq}\ y{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -47,8 +60,14 @@ We can now prove simple theorems in this abstract setting, for example that \isa{{\isacharless}{\isacharless}} is not symmetric:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcolon}{\isacharcolon}{\isacharprime}a{\isacharcolon}{\isacharcolon}parord{\isacharparenright}\ {\isacharless}{\isacharless}\ y\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isasymnot}\ y\ {\isacharless}{\isacharless}\ x{\isacharparenright}\ {\isacharequal}\ True{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcolon}{\isacharcolon}{\isacharprime}a{\isacharcolon}{\isacharcolon}parord{\isacharparenright}\ {\isacharless}{\isacharless}\ y\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isasymnot}\ y\ {\isacharless}{\isacharless}\ x{\isacharparenright}\ {\isacharequal}\ True{\isachardoublequote}\isamarkupfalse% % \begin{isamarkuptxt}% \noindent @@ -63,8 +82,15 @@ \isa{{\isacharprime}a{\isacharcolon}{\isacharcolon}ordrel} (as required in the type of \isa{{\isacharless}{\isacharless}}), when the proposition is not a theorem. The proof is easy:% \end{isamarkuptxt}% +\isamarkupfalse% +\isacommand{by}{\isacharparenleft}simp\ add{\isacharcolon}\ less{\isacharunderscore}le\ antisym{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{by}{\isacharparenleft}simp\ add{\isacharcolon}\ less{\isacharunderscore}le\ antisym{\isacharparenright}\isamarkupfalse% % \begin{isamarkuptext}% We could now continue in this vein and develop a whole theory of @@ -73,10 +99,16 @@ prove that the types in question, for example \isa{bool}, are indeed instances of \isa{parord}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{instance}\ bool\ {\isacharcolon}{\isacharcolon}\ parord\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}\ intro{\isacharunderscore}classes\isamarkupfalse% +\isacommand{apply}\ intro{\isacharunderscore}classes\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -92,10 +124,17 @@ once we have unfolded the definitions of \isa{{\isacharless}{\isacharless}} and \isa{{\isacharless}{\isacharless}{\isacharequal}} at type \isa{bool}:% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}simp{\isacharunderscore}all\ {\isacharparenleft}no{\isacharunderscore}asm{\isacharunderscore}use{\isacharparenright}\ only{\isacharcolon}\ le{\isacharunderscore}bool{\isacharunderscore}def\ less{\isacharunderscore}bool{\isacharunderscore}def{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{by}{\isacharparenleft}blast{\isacharcomma}\ blast{\isacharcomma}\ blast{\isacharcomma}\ blast{\isacharparenright}\isamarkupfalse% +\isacommand{by}{\isacharparenleft}blast{\isacharcomma}\ blast{\isacharcomma}\ blast{\isacharcomma}\ blast{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -103,10 +142,23 @@ We can now apply our single lemma above in the context of booleans:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}P{\isacharcolon}{\isacharcolon}bool{\isacharparenright}\ {\isacharless}{\isacharless}\ Q\ {\isasymLongrightarrow}\ {\isasymnot}{\isacharparenleft}Q\ {\isacharless}{\isacharless}\ P{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +\isacommand{by}\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -125,9 +177,9 @@ If any two elements of a partial order are comparable it is a \textbf{linear} or \textbf{total} order:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{axclass}\ linord\ {\isacharless}\ parord\isanewline -linear{\isacharcolon}\ {\isachardoublequote}x\ {\isacharless}{\isacharless}{\isacharequal}\ y\ {\isasymor}\ y\ {\isacharless}{\isacharless}{\isacharequal}\ x{\isachardoublequote}\isamarkupfalse% +linear{\isacharcolon}\ {\isachardoublequote}x\ {\isacharless}{\isacharless}{\isacharequal}\ y\ {\isasymor}\ y\ {\isacharless}{\isacharless}{\isacharequal}\ x{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -135,8 +187,14 @@ Therefore we can show that linearity can be expressed in terms of \isa{{\isacharless}{\isacharless}} as follows:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isasymAnd}x{\isacharcolon}{\isacharcolon}{\isacharprime}a{\isacharcolon}{\isacharcolon}linord{\isachardot}\ x\ {\isacharless}{\isacharless}\ y\ {\isasymor}\ x\ {\isacharequal}\ y\ {\isasymor}\ y\ {\isacharless}{\isacharless}\ x{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}{\isacharparenleft}simp\ add{\isacharcolon}\ less{\isacharunderscore}le{\isacharparenright}\isanewline \isamarkupfalse% @@ -144,7 +202,14 @@ \isamarkupfalse% \isacommand{apply}\ blast\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% Linear orders are an example of subclassing\index{subclasses} @@ -163,11 +228,11 @@ An alternative axiomatization of partial orders takes \isa{{\isacharless}{\isacharless}} rather than \isa{{\isacharless}{\isacharless}{\isacharequal}} as the primary concept. The result is a \textbf{strict} order:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{axclass}\ strord\ {\isacharless}\ ordrel\isanewline irrefl{\isacharcolon}\ \ \ \ \ {\isachardoublequote}{\isasymnot}\ x\ {\isacharless}{\isacharless}\ x{\isachardoublequote}\isanewline less{\isacharunderscore}trans{\isacharcolon}\ {\isachardoublequote}{\isasymlbrakk}\ x\ {\isacharless}{\isacharless}\ y{\isacharsemicolon}\ y\ {\isacharless}{\isacharless}\ z\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ x\ {\isacharless}{\isacharless}\ z{\isachardoublequote}\isanewline -le{\isacharunderscore}less{\isacharcolon}\ \ \ \ {\isachardoublequote}x\ {\isacharless}{\isacharless}{\isacharequal}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}{\isacharless}\ y\ {\isasymor}\ x\ {\isacharequal}\ y{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +le{\isacharunderscore}less{\isacharcolon}\ \ \ \ {\isachardoublequote}x\ {\isacharless}{\isacharless}{\isacharequal}\ y\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharless}{\isacharless}\ y\ {\isasymor}\ x\ {\isacharequal}\ y{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -175,10 +240,16 @@ prove one direction, namely that partial orders are a subclass of strict orders.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{instance}\ parord\ {\isacharless}\ strord\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}\ intro{\isacharunderscore}classes\isamarkupfalse% +\isacommand{apply}\ intro{\isacharunderscore}classes\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -192,14 +263,21 @@ Assuming \isa{{\isacharprime}a\ {\isacharcolon}{\isacharcolon}\ parord}, the three axioms of class \isa{strord} are easily proved:% \end{isamarkuptxt}% -\ \ \isamarkuptrue% +\ \ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}simp{\isacharunderscore}all\ {\isacharparenleft}no{\isacharunderscore}asm{\isacharunderscore}use{\isacharparenright}\ add{\isacharcolon}\ less{\isacharunderscore}le{\isacharparenright}\isanewline \ \isamarkupfalse% \isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ trans\ antisym{\isacharparenright}\isanewline \isamarkupfalse% \isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ refl{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% The subclass relation must always be acyclic. Therefore Isabelle will @@ -216,12 +294,12 @@ \bfindex{multiple inheritance}. For example, we could define the classes of well-founded orderings and well-orderings:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{axclass}\ wford\ {\isacharless}\ parord\isanewline wford{\isacharcolon}\ {\isachardoublequote}wf\ {\isacharbraceleft}{\isacharparenleft}y{\isacharcomma}x{\isacharparenright}{\isachardot}\ y\ {\isacharless}{\isacharless}\ x{\isacharbraceright}{\isachardoublequote}\isanewline \isanewline \isamarkupfalse% -\isacommand{axclass}\ wellord\ {\isacharless}\ linord{\isacharcomma}\ wford\isamarkupfalse% +\isacommand{axclass}\ wellord\ {\isacharless}\ linord{\isacharcomma}\ wford\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -289,8 +367,19 @@ classes readily become inconsistent in practice. Now we know this need not worry us.% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Types/document/Numbers.tex --- a/doc-src/TutorialI/Types/document/Numbers.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Types/document/Numbers.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,33 +1,73 @@ % \begin{isabellebody}% \def\isabellecontext{Numbers}% +% +\isadelimtheory \isanewline -\isacommand{theory}\ Numbers\ \isakeyword{imports}\ Real\ \isakeyword{begin}\isanewline +% +\endisadelimtheory +% +\isatagtheory +\isamarkupfalse% +\isacommand{theory}\ Numbers\ \isakeyword{imports}\ Real\ \isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory \isanewline +% +\endisadelimtheory +% +\isadelimML +\isanewline +% +\endisadelimML +% +\isatagML \isamarkupfalse% \isacommand{ML}\ {\isachardoublequote}Pretty{\isachardot}setmargin\ {\isadigit{6}}{\isadigit{4}}{\isachardoublequote}\isanewline \isamarkupfalse% -\isacommand{ML}\ {\isachardoublequote}IsarOutput{\isachardot}indent\ {\isacharcolon}{\isacharequal}\ {\isadigit{0}}{\isachardoublequote}\isamarkupfalse% +\isacommand{ML}\ {\isachardoublequote}IsarOutput{\isachardot}indent\ {\isacharcolon}{\isacharequal}\ {\isadigit{0}}{\isachardoublequote}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +\isamarkuptrue% % \begin{isamarkuptext}% numeric literals; default simprules; can re-orient% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}{\isadigit{2}}\ {\isacharasterisk}\ m\ {\isacharequal}\ m\ {\isacharplus}\ m{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{lemma}\ {\isachardoublequote}{\isadigit{2}}\ {\isacharasterisk}\ m\ {\isacharequal}\ m\ {\isacharplus}\ m{\isachardoublequote}\isamarkupfalse% % \begin{isamarkuptxt}% \begin{isabelle}% \ {\isadigit{1}}{\isachardot}\ {\isacharparenleft}{\isadigit{2}}{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharasterisk}\ m\ {\isacharequal}\ m\ {\isacharplus}\ m% \end{isabelle}% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{oops}\isanewline +\isamarkupfalse% +\isacommand{oops}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline \isanewline \isamarkupfalse% \isacommand{consts}\ h\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}nat\ {\isasymRightarrow}\ nat{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{recdef}\ h\ {\isachardoublequote}{\isacharbraceleft}{\isacharbraceright}{\isachardoublequote}\isanewline -{\isachardoublequote}h\ i\ {\isacharequal}\ {\isacharparenleft}if\ i\ {\isacharequal}\ {\isadigit{3}}\ then\ {\isadigit{2}}\ else\ i{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +{\isachardoublequote}h\ i\ {\isacharequal}\ {\isacharparenleft}if\ i\ {\isacharequal}\ {\isadigit{3}}\ then\ {\isadigit{2}}\ else\ i{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \isa{h\ {\isadigit{3}}\ {\isacharequal}\ {\isadigit{2}}} @@ -73,16 +113,22 @@ these form add_ac; similarly there is mult_ac% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}Suc{\isacharparenleft}i\ {\isacharplus}\ j{\isacharasterisk}l{\isacharasterisk}k\ {\isacharplus}\ m{\isacharasterisk}n{\isacharparenright}\ {\isacharequal}\ f\ {\isacharparenleft}n{\isacharasterisk}m\ {\isacharplus}\ i\ {\isacharplus}\ k{\isacharasterisk}j{\isacharasterisk}l{\isacharparenright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{lemma}\ {\isachardoublequote}Suc{\isacharparenleft}i\ {\isacharplus}\ j{\isacharasterisk}l{\isacharasterisk}k\ {\isacharplus}\ m{\isacharasterisk}n{\isacharparenright}\ {\isacharequal}\ f\ {\isacharparenleft}n{\isacharasterisk}m\ {\isacharplus}\ i\ {\isacharplus}\ k{\isacharasterisk}j{\isacharasterisk}l{\isacharparenright}{\isachardoublequote}\isamarkupfalse% % \begin{isamarkuptxt}% \begin{isabelle}% \ {\isadigit{1}}{\isachardot}\ Suc\ {\isacharparenleft}i\ {\isacharplus}\ j\ {\isacharasterisk}\ l\ {\isacharasterisk}\ k\ {\isacharplus}\ m\ {\isacharasterisk}\ n{\isacharparenright}\ {\isacharequal}\ f\ {\isacharparenleft}n\ {\isacharasterisk}\ m\ {\isacharplus}\ i\ {\isacharplus}\ k\ {\isacharasterisk}\ j\ {\isacharasterisk}\ l{\isacharparenright}% \end{isabelle}% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\ {\isacharparenleft}simp\ add{\isacharcolon}\ add{\isacharunderscore}ac\ mult{\isacharunderscore}ac{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}\ {\isacharparenleft}simp\ add{\isacharcolon}\ add{\isacharunderscore}ac\ mult{\isacharunderscore}ac{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -90,8 +136,15 @@ \isaindent{\ {\isadigit{1}}{\isachardot}\ }f\ {\isacharparenleft}i\ {\isacharplus}\ {\isacharparenleft}m\ {\isacharasterisk}\ n\ {\isacharplus}\ j\ {\isacharasterisk}\ {\isacharparenleft}k\ {\isacharasterisk}\ l{\isacharparenright}{\isacharparenright}{\isacharparenright}% \end{isabelle}% \end{isamarkuptxt}% +\isamarkupfalse% +\isacommand{oops}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{oops}\isamarkupfalse% % \begin{isamarkuptext}% \begin{isabelle}% @@ -114,8 +167,14 @@ \end{isabelle} \rulename{nat_diff_split}% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}n\ {\isacharminus}\ {\isadigit{1}}{\isacharparenright}\ {\isacharasterisk}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ n\ {\isacharasterisk}\ n\ {\isacharminus}\ {\isacharparenleft}{\isadigit{1}}{\isacharcolon}{\isacharcolon}nat{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}\ {\isacharparenleft}clarsimp\ split{\isacharcolon}\ nat{\isacharunderscore}diff{\isacharunderscore}split\ iff\ del{\isacharcolon}\ less{\isacharunderscore}Suc{\isadigit{0}}{\isacharparenright}\isanewline \ % @@ -127,11 +186,24 @@ \isamarkupfalse% \isacommand{apply}\ {\isacharparenleft}subgoal{\isacharunderscore}tac\ {\isachardoublequote}n{\isacharequal}{\isadigit{0}}{\isachardoublequote}{\isacharcomma}\ force{\isacharcomma}\ arith{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isanewline +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof \isanewline \isanewline \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}n\ {\isacharminus}\ {\isadigit{2}}{\isacharparenright}\ {\isacharasterisk}\ {\isacharparenleft}n\ {\isacharplus}\ {\isadigit{2}}{\isacharparenright}\ {\isacharequal}\ n\ {\isacharasterisk}\ n\ {\isacharminus}\ {\isacharparenleft}{\isadigit{4}}{\isacharcolon}{\isacharcolon}nat{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% \isacommand{apply}\ {\isacharparenleft}simp\ split{\isacharcolon}\ nat{\isacharunderscore}diff{\isacharunderscore}split{\isacharcomma}\ clarify{\isacharparenright}\isanewline \ % @@ -143,7 +215,14 @@ \isamarkupfalse% \isacommand{apply}\ {\isacharparenleft}subgoal{\isacharunderscore}tac\ {\isachardoublequote}n{\isacharequal}{\isadigit{0}}\ {\isacharbar}\ n{\isacharequal}{\isadigit{1}}{\isachardoublequote}{\isacharcomma}\ force{\isacharcomma}\ arith{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \begin{isabelle}% @@ -261,15 +340,41 @@ \end{isabelle} \rulename{zmod_zmult2_eq}% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}abs\ {\isacharparenleft}x{\isacharplus}y{\isacharparenright}\ {\isasymle}\ abs\ x\ {\isacharplus}\ abs\ {\isacharparenleft}y\ {\isacharcolon}{\isacharcolon}\ int{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}\ arith\isanewline +\isacommand{by}\ arith% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof \isanewline \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}abs\ {\isacharparenleft}{\isadigit{2}}{\isacharasterisk}x{\isacharparenright}\ {\isacharequal}\ {\isadigit{2}}\ {\isacharasterisk}\ abs\ {\isacharparenleft}x\ {\isacharcolon}{\isacharcolon}\ int{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ add{\isacharcolon}\ abs{\isacharunderscore}if{\isacharparenright}\isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}simp\ add{\isacharcolon}\ abs{\isacharunderscore}if{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% Induction rules for the Integers @@ -341,48 +446,87 @@ \end{isabelle} \rulename{add_divide_distrib}% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isadigit{3}}{\isacharslash}{\isadigit{4}}\ {\isacharless}\ {\isacharparenleft}{\isadigit{7}}{\isacharslash}{\isadigit{8}}\ {\isacharcolon}{\isacharcolon}\ real{\isacharparenright}{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}\ simp\ \isanewline +\isacommand{by}\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\ \isanewline +% +\endisadelimproof \isanewline \isamarkupfalse% -\isacommand{lemma}\ {\isachardoublequote}P\ {\isacharparenleft}{\isacharparenleft}{\isadigit{3}}{\isacharslash}{\isadigit{4}}{\isacharparenright}\ {\isacharasterisk}\ {\isacharparenleft}{\isadigit{8}}{\isacharslash}{\isadigit{1}}{\isadigit{5}}\ {\isacharcolon}{\isacharcolon}\ real{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}P\ {\isacharparenleft}{\isacharparenleft}{\isadigit{3}}{\isacharslash}{\isadigit{4}}{\isacharparenright}\ {\isacharasterisk}\ {\isacharparenleft}{\isadigit{8}}{\isacharslash}{\isadigit{1}}{\isadigit{5}}\ {\isacharcolon}{\isacharcolon}\ real{\isacharparenright}{\isacharparenright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% \ {\isadigit{1}}{\isachardot}\ P\ {\isacharparenleft}{\isadigit{3}}\ {\isacharslash}\ {\isadigit{4}}\ {\isacharasterisk}\ {\isacharparenleft}{\isadigit{8}}\ {\isacharslash}\ {\isadigit{1}}{\isadigit{5}}{\isacharparenright}{\isacharparenright}% \end{isabelle}% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\ simp\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}\ simp\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% \ {\isadigit{1}}{\isachardot}\ P\ {\isacharparenleft}{\isadigit{2}}\ {\isacharslash}\ {\isadigit{5}}{\isacharparenright}% \end{isabelle}% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{oops}\isanewline +\isamarkupfalse% +\isacommand{oops}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isanewline \isanewline \isamarkupfalse% -\isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}{\isadigit{3}}{\isacharslash}{\isadigit{4}}{\isacharparenright}\ {\isacharasterisk}\ {\isacharparenleft}{\isadigit{8}}{\isacharslash}{\isadigit{1}}{\isadigit{5}}{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x\ {\isacharcolon}{\isacharcolon}\ real{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}{\isadigit{3}}{\isacharslash}{\isadigit{4}}{\isacharparenright}\ {\isacharasterisk}\ {\isacharparenleft}{\isadigit{8}}{\isacharslash}{\isadigit{1}}{\isadigit{5}}{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x\ {\isacharcolon}{\isacharcolon}\ real{\isacharparenright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% \ {\isadigit{1}}{\isachardot}\ {\isadigit{3}}\ {\isacharslash}\ {\isadigit{4}}\ {\isacharasterisk}\ {\isacharparenleft}{\isadigit{8}}\ {\isacharslash}\ {\isadigit{1}}{\isadigit{5}}{\isacharparenright}\ {\isacharless}\ x% \end{isabelle}% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}\ simp\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}\ simp\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% \ {\isadigit{1}}{\isachardot}\ {\isadigit{2}}\ {\isacharless}\ x\ {\isacharasterisk}\ {\isadigit{5}}% \end{isabelle}% \end{isamarkuptxt}% +\isamarkupfalse% +\isacommand{oops}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{oops}\isamarkupfalse% % \begin{isamarkuptext}% Ring and Field @@ -409,8 +553,21 @@ \end{isabelle} \rulename{field_mult_cancel_right}% \end{isamarkuptext}% +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isamarkupfalse% +\isacommand{ML}{\isacharbraceleft}{\isacharasterisk}set\ show{\isacharunderscore}sorts{\isacharasterisk}{\isacharbraceright}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML \isamarkuptrue% -\isacommand{ML}{\isacharbraceleft}{\isacharasterisk}set\ show{\isacharunderscore}sorts{\isacharasterisk}{\isacharbraceright}\isamarkupfalse% % \begin{isamarkuptext}% effect of show sorts on the above @@ -421,8 +578,21 @@ \end{isabelle} \rulename{field_mult_cancel_right}% \end{isamarkuptext}% +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isamarkupfalse% +\isacommand{ML}{\isacharbraceleft}{\isacharasterisk}reset\ show{\isacharunderscore}sorts{\isacharasterisk}{\isacharbraceright}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML \isamarkuptrue% -\isacommand{ML}{\isacharbraceleft}{\isacharasterisk}reset\ show{\isacharunderscore}sorts{\isacharasterisk}{\isacharbraceright}\isamarkupfalse% % \begin{isamarkuptext}% absolute value @@ -457,9 +627,21 @@ \end{isabelle} \rulename{power_abs}% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{end}\isanewline +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory \isamarkupfalse% +\isacommand{end}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Types/document/Overloading.tex --- a/doc-src/TutorialI/Types/document/Overloading.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Types/document/Overloading.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,10 +1,36 @@ % \begin{isabellebody}% \def\isabellecontext{Overloading}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \isamarkupfalse% \isacommand{instance}\ list\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}type{\isacharparenright}ordrel\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}\ intro{\isacharunderscore}classes\isamarkupfalse% +\isacommand{by}\ intro{\isacharunderscore}classes% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -15,13 +41,24 @@ Of course we should also define the meaning of \isa{{\isacharless}{\isacharless}{\isacharequal}} and \isa{{\isacharless}{\isacharless}} on lists:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{defs}\ {\isacharparenleft}\isakeyword{overloaded}{\isacharparenright}\isanewline prefix{\isacharunderscore}def{\isacharcolon}\isanewline \ \ {\isachardoublequote}xs\ {\isacharless}{\isacharless}{\isacharequal}\ {\isacharparenleft}ys{\isacharcolon}{\isacharcolon}{\isacharprime}a{\isacharcolon}{\isacharcolon}ordrel\ list{\isacharparenright}\ \ {\isasymequiv}\ \ {\isasymexists}zs{\isachardot}\ ys\ {\isacharequal}\ xs{\isacharat}zs{\isachardoublequote}\isanewline strict{\isacharunderscore}prefix{\isacharunderscore}def{\isacharcolon}\isanewline -\ \ {\isachardoublequote}xs\ {\isacharless}{\isacharless}\ {\isacharparenleft}ys{\isacharcolon}{\isacharcolon}{\isacharprime}a{\isacharcolon}{\isacharcolon}ordrel\ list{\isacharparenright}\ \ \ {\isasymequiv}\ \ xs\ {\isacharless}{\isacharless}{\isacharequal}\ ys\ {\isasymand}\ xs\ {\isasymnoteq}\ ys{\isachardoublequote}\isamarkupfalse% -\isamarkupfalse% +\ \ {\isachardoublequote}xs\ {\isacharless}{\isacharless}\ {\isacharparenleft}ys{\isacharcolon}{\isacharcolon}{\isacharprime}a{\isacharcolon}{\isacharcolon}ordrel\ list{\isacharparenright}\ \ \ {\isasymequiv}\ \ xs\ {\isacharless}{\isacharless}{\isacharequal}\ ys\ {\isasymand}\ xs\ {\isasymnoteq}\ ys{\isachardoublequote}% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Types/document/Overloading0.tex --- a/doc-src/TutorialI/Types/document/Overloading0.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Types/document/Overloading0.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Overloading{\isadigit{0}}}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% We start with a concept that is required for type classes but already @@ -18,18 +31,18 @@ If we want to introduce the notion of an \emph{inverse} for arbitrary types we give it a polymorphic type% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{consts}\ inverse\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\isamarkupfalse% +\isamarkupfalse% +\isacommand{consts}\ inverse\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent and provide different definitions at different instances:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{defs}\ {\isacharparenleft}\isakeyword{overloaded}{\isacharparenright}\isanewline inverse{\isacharunderscore}bool{\isacharcolon}\ {\isachardoublequote}inverse{\isacharparenleft}b{\isacharcolon}{\isacharcolon}bool{\isacharparenright}\ {\isasymequiv}\ {\isasymnot}\ b{\isachardoublequote}\isanewline inverse{\isacharunderscore}set{\isacharcolon}\ \ {\isachardoublequote}inverse{\isacharparenleft}A{\isacharcolon}{\isacharcolon}{\isacharprime}a\ set{\isacharparenright}\ {\isasymequiv}\ {\isacharminus}A{\isachardoublequote}\isanewline -inverse{\isacharunderscore}pair{\isacharcolon}\ {\isachardoublequote}inverse{\isacharparenleft}p{\isacharparenright}\ {\isasymequiv}\ {\isacharparenleft}inverse{\isacharparenleft}fst\ p{\isacharparenright}{\isacharcomma}\ inverse{\isacharparenleft}snd\ p{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +inverse{\isacharunderscore}pair{\isacharcolon}\ {\isachardoublequote}inverse{\isacharparenleft}p{\isacharparenright}\ {\isasymequiv}\ {\isacharparenleft}inverse{\isacharparenleft}fst\ p{\isacharparenright}{\isacharcomma}\ inverse{\isacharparenleft}snd\ p{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -48,8 +61,19 @@ unspecified constants does not endanger soundness, but it is pointless. To prevent such terms from even being formed requires the use of type classes.% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Types/document/Overloading1.tex --- a/doc-src/TutorialI/Types/document/Overloading1.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Types/document/Overloading1.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Overloading{\isadigit{1}}}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsubsubsection{Controlled Overloading with Type Classes% } @@ -13,8 +26,8 @@ to avoid clashes with \isa{{\isacharless}} and \isa{{\isacharless}{\isacharequal}} in theory \isa{Main}. To restrict the application of \isa{{\isacharless}{\isacharless}} and \isa{{\isacharless}{\isacharless}{\isacharequal}} we introduce the class \isa{ordrel}:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{axclass}\ ordrel\ {\isacharless}\ type\isamarkupfalse% +\isamarkupfalse% +\isacommand{axclass}\ ordrel\ {\isacharless}\ type\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -25,9 +38,9 @@ Its sole purpose is to restrict the use of overloaded constants to meaningful instances:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{consts}\ less\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcolon}{\isacharcolon}ordrel{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ bool{\isachardoublequote}\ \ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isacharless}{\isacharless}{\isachardoublequote}\ \ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\isanewline -\ \ \ \ \ \ \ le\ \ \ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcolon}{\isacharcolon}ordrel{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ bool{\isachardoublequote}\ \ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isacharless}{\isacharless}{\isacharequal}{\isachardoublequote}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\isamarkupfalse% +\ \ \ \ \ \ \ le\ \ \ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcolon}{\isacharcolon}ordrel{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ bool{\isachardoublequote}\ \ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequote}{\isacharless}{\isacharless}{\isacharequal}{\isachardoublequote}\ {\isadigit{5}}{\isadigit{0}}{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -39,8 +52,14 @@ into \isa{ordrel} we need to declare a type to be an \bfindex{instance} of \isa{ordrel}:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{instance}\ bool\ {\isacharcolon}{\isacharcolon}\ ordrel% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{instance}\ bool\ {\isacharcolon}{\isacharcolon}\ ordrel\isamarkupfalse% % \begin{isamarkuptxt}% \noindent @@ -49,8 +68,15 @@ There are none, but we still need to finish that proof, which we do by invoking the \methdx{intro_classes} method:% \end{isamarkuptxt}% +\isamarkupfalse% +\isacommand{by}\ intro{\isacharunderscore}classes% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkuptrue% -\isacommand{by}\ intro{\isacharunderscore}classes\isamarkupfalse% % \begin{isamarkuptext}% \noindent @@ -60,19 +86,32 @@ Although terms like \isa{False\ {\isacharless}{\isacharless}{\isacharequal}\ P} are now legal, we still need to say what the relation symbols actually mean at type \isa{bool}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{defs}\ {\isacharparenleft}\isakeyword{overloaded}{\isacharparenright}\isanewline le{\isacharunderscore}bool{\isacharunderscore}def{\isacharcolon}\ \ {\isachardoublequote}P\ {\isacharless}{\isacharless}{\isacharequal}\ Q\ {\isasymequiv}\ P\ {\isasymlongrightarrow}\ Q{\isachardoublequote}\isanewline -less{\isacharunderscore}bool{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequote}P\ {\isacharless}{\isacharless}\ Q\ {\isasymequiv}\ {\isasymnot}P\ {\isasymand}\ Q{\isachardoublequote}\isamarkupfalse% +less{\isacharunderscore}bool{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequote}P\ {\isacharless}{\isacharless}\ Q\ {\isasymequiv}\ {\isasymnot}P\ {\isasymand}\ Q{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent Now \isa{False\ {\isacharless}{\isacharless}{\isacharequal}\ P} is provable:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}False\ {\isacharless}{\isacharless}{\isacharequal}\ P{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}{\isacharparenleft}simp\ add{\isacharcolon}\ le{\isacharunderscore}bool{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% +\isacommand{by}{\isacharparenleft}simp\ add{\isacharcolon}\ le{\isacharunderscore}bool{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -80,8 +119,19 @@ To make it well-typed, we need to make lists a type of class \isa{ordrel}:% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Types/document/Overloading2.tex --- a/doc-src/TutorialI/Types/document/Overloading2.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Types/document/Overloading2.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Overloading{\isadigit{2}}}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \begin{isamarkuptext}% Of course this is not the only possible definition of the two relations. @@ -9,15 +22,28 @@ the elements of the list must also be of class \isa{ordrel} to permit their comparison:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{instance}\ list\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}ordrel{\isacharparenright}ordrel\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}\ intro{\isacharunderscore}classes\isanewline +\isacommand{by}\ intro{\isacharunderscore}classes% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +\isanewline +% +\endisadelimproof \isanewline \isamarkupfalse% \isacommand{defs}\ {\isacharparenleft}\isakeyword{overloaded}{\isacharparenright}\isanewline le{\isacharunderscore}list{\isacharunderscore}def{\isacharcolon}\ {\isachardoublequote}xs\ {\isacharless}{\isacharless}{\isacharequal}\ {\isacharparenleft}ys{\isacharcolon}{\isacharcolon}{\isacharprime}a{\isacharcolon}{\isacharcolon}ordrel\ list{\isacharparenright}\ {\isasymequiv}\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ size\ xs\ {\isacharequal}\ size\ ys\ {\isasymand}\ {\isacharparenleft}{\isasymforall}i{\isacharless}size\ xs{\isachardot}\ xs{\isacharbang}i\ {\isacharless}{\isacharless}{\isacharequal}\ ys{\isacharbang}i{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ \ \ \ \ \ \ \ \ \ \ \ \ size\ xs\ {\isacharequal}\ size\ ys\ {\isasymand}\ {\isacharparenleft}{\isasymforall}i{\isacharless}size\ xs{\isachardot}\ xs{\isacharbang}i\ {\isacharless}{\isacharless}{\isacharequal}\ ys{\isacharbang}i{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -50,8 +76,19 @@ \end{center} And analogously for \isa{{\isacharless}} instead of \isa{{\isasymle}}.% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Types/document/Pairs.tex --- a/doc-src/TutorialI/Types/document/Pairs.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Types/document/Pairs.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Pairs}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsection{Pairs and Tuples% } @@ -64,10 +77,23 @@ \begin{isamarkuptext}% The most obvious approach is the brute force expansion of \isa{split}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}{\isasymlambda}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}{\isachardot}x{\isacharparenright}\ p\ {\isacharequal}\ fst\ p{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}{\isacharparenleft}simp\ add{\isacharcolon}\ split{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% +\isacommand{by}{\isacharparenleft}simp\ add{\isacharcolon}\ split{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% This works well if rewriting with \isa{split{\isacharunderscore}def} finishes the @@ -87,10 +113,16 @@ rule \isa{split{\isacharunderscore}split} replaces \isa{p} by a pair:% \index{*split (method)}% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}{\isasymlambda}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}{\isachardot}y{\isacharparenright}\ p\ {\isacharequal}\ snd\ p{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}split\ split{\isacharunderscore}split{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}split\ split{\isacharunderscore}split{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -99,18 +131,42 @@ This subgoal is easily proved by simplification. Thus we could have combined simplification and splitting in one command that proves the goal outright:% \end{isamarkuptxt}% -\isamarkuptrue% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isamarkupfalse% -\isacommand{by}{\isacharparenleft}simp\ split{\isacharcolon}\ split{\isacharunderscore}split{\isacharparenright}\isamarkupfalse% +\isacommand{by}{\isacharparenleft}simp\ split{\isacharcolon}\ split{\isacharunderscore}split{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% Let us look at a second example:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}let\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isacharequal}\ p\ in\ fst\ p\ {\isacharequal}\ x{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp\ only{\isacharcolon}\ Let{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}simp\ only{\isacharcolon}\ Let{\isacharunderscore}def{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -119,11 +175,23 @@ A paired \isa{let} reduces to a paired $\lambda$-abstraction, which can be split as above. The same is true for paired set comprehension:% \end{isamarkuptxt}% -\isamarkuptrue% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}p\ {\isasymin}\ {\isacharbraceleft}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}{\isachardot}\ x{\isacharequal}y{\isacharbraceright}\ {\isasymlongrightarrow}\ fst\ p\ {\isacharequal}\ snd\ p{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}\ simp\isamarkupfalse% +\isacommand{apply}\ simp\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -134,9 +202,21 @@ \isa{{\isasymlambda}{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}{\isachardot}\ x\ {\isacharequal}\ y} is short for \isa{{\isasymlambda}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}{\isachardot}\ x{\isacharequal}y}. The same proof procedure works for% \end{isamarkuptxt}% -\isamarkuptrue% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkupfalse% -\isacommand{lemma}\ {\isachardoublequote}p\ {\isasymin}\ {\isacharbraceleft}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}{\isachardot}\ x{\isacharequal}y{\isacharbraceright}\ {\isasymLongrightarrow}\ fst\ p\ {\isacharequal}\ snd\ p{\isachardoublequote}\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}p\ {\isasymin}\ {\isacharbraceleft}{\isacharparenleft}x{\isacharcomma}y{\isacharparenright}{\isachardot}\ x{\isacharequal}y{\isacharbraceright}\ {\isasymLongrightarrow}\ fst\ p\ {\isacharequal}\ snd\ p{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof +\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -146,20 +226,32 @@ However, splitting \isa{split} is not always a solution, as no \isa{split} may be present in the goal. Consider the following function:% \end{isamarkuptxt}% -\isamarkuptrue% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkupfalse% \isacommand{consts}\ swap\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b\ {\isasymRightarrow}\ {\isacharprime}b\ {\isasymtimes}\ {\isacharprime}a{\isachardoublequote}\isanewline \isamarkupfalse% \isacommand{primrec}\isanewline -\ \ {\isachardoublequote}swap\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}y{\isacharcomma}x{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}swap\ {\isacharparenleft}x{\isacharcomma}y{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}y{\isacharcomma}x{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent Note that the above \isacommand{primrec} definition is admissible because \isa{{\isasymtimes}} is a datatype. When we now try to prove% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}swap{\isacharparenleft}swap\ p{\isacharparenright}\ {\isacharequal}\ p{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{lemma}\ {\isachardoublequote}swap{\isacharparenleft}swap\ p{\isacharparenright}\ {\isacharequal}\ p{\isachardoublequote}\isamarkupfalse% % \begin{isamarkuptxt}% \noindent @@ -168,8 +260,8 @@ time there is no \isa{split} in sight. In this case the only thing we can do is to split the term by hand:% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}{\isacharparenleft}case{\isacharunderscore}tac\ p{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}case{\isacharunderscore}tac\ p{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -187,11 +279,23 @@ You can split \emph{all} \isa{{\isasymAnd}}-quantified variables in a goal with the rewrite rule \isa{split{\isacharunderscore}paired{\isacharunderscore}all}:% \end{isamarkuptxt}% -\isamarkuptrue% +% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof \isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isasymAnd}p\ q{\isachardot}\ swap{\isacharparenleft}swap\ p{\isacharparenright}\ {\isacharequal}\ q\ {\isasymlongrightarrow}\ p\ {\isacharequal}\ q{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp\ only{\isacharcolon}\ split{\isacharunderscore}paired{\isacharunderscore}all{\isacharparenright}\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}simp\ only{\isacharcolon}\ split{\isacharunderscore}paired{\isacharunderscore}all{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \noindent @@ -201,10 +305,17 @@ \isaindent{\ {\isadigit{1}}{\isachardot}\ \ \ \ }{\isacharparenleft}a{\isacharcomma}\ b{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}aa{\isacharcomma}\ ba{\isacharparenright}% \end{isabelle}% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}\ simp\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -217,20 +328,44 @@ The following command could fail (here it does not) where two separate \isa{simp} applications succeed.% \end{isamarkuptext}% -\isamarkuptrue% +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{apply}{\isacharparenleft}simp\ add{\isacharcolon}\ split{\isacharunderscore}paired{\isacharunderscore}all{\isacharparenright}\isamarkupfalse% -\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}simp\ add{\isacharcolon}\ split{\isacharunderscore}paired{\isacharunderscore}all{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent Finally, the simplifier automatically splits all \isa{{\isasymforall}} and \isa{{\isasymexists}}-quantified variables:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isasymforall}p{\isachardot}\ {\isasymexists}q{\isachardot}\ swap\ p\ {\isacharequal}\ swap\ q{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +\isacommand{by}\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -245,8 +380,19 @@ (\isa{split{\isacharunderscore}paired{\isacharunderscore}Ex}) \end{center}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Types/document/Records.tex --- a/doc-src/TutorialI/Types/document/Records.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Types/document/Records.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,11 +1,24 @@ % \begin{isabellebody}% \def\isabellecontext{Records}% +\isamarkuptrue% % \isamarkupheader{Records \label{sec:records}% } +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \isamarkuptrue% -\isamarkupfalse% % \begin{isamarkuptext}% \index{records|(}% @@ -43,20 +56,20 @@ fields, which are packaged internally to hold up the perception of the record as a distinguished entity. Here is a simple example:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{record}\ point\ {\isacharequal}\isanewline \ \ Xcoord\ {\isacharcolon}{\isacharcolon}\ int\isanewline -\ \ Ycoord\ {\isacharcolon}{\isacharcolon}\ int\isamarkupfalse% +\ \ Ycoord\ {\isacharcolon}{\isacharcolon}\ int\isamarkuptrue% % \begin{isamarkuptext}% Records of type \isa{point} have two fields named \isa{Xcoord} and \isa{Ycoord}, both of type~\isa{int}. We now define a constant of type \isa{point}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{constdefs}\isanewline \ \ pt{\isadigit{1}}\ {\isacharcolon}{\isacharcolon}\ point\isanewline -\ \ {\isachardoublequote}pt{\isadigit{1}}\ {\isasymequiv}\ {\isacharparenleft}{\isacharbar}\ Xcoord\ {\isacharequal}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharcomma}\ Ycoord\ {\isacharequal}\ {\isadigit{2}}{\isadigit{3}}\ {\isacharbar}{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}pt{\isadigit{1}}\ {\isasymequiv}\ {\isacharparenleft}{\isacharbar}\ Xcoord\ {\isacharequal}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharcomma}\ Ycoord\ {\isacharequal}\ {\isadigit{2}}{\isadigit{3}}\ {\isacharbar}{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% We see above the ASCII notation for record brackets. You can also @@ -64,31 +77,57 @@ expressions can be also written directly with individual fields. The type name above is merely an abbreviation.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{constdefs}\isanewline \ \ pt{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharcolon}\ int{\isacharcomma}\ Ycoord\ {\isacharcolon}{\isacharcolon}\ int{\isasymrparr}{\isachardoublequote}\isanewline -\ \ {\isachardoublequote}pt{\isadigit{2}}\ {\isasymequiv}\ {\isasymlparr}Xcoord\ {\isacharequal}\ {\isacharminus}{\isadigit{4}}{\isadigit{5}}{\isacharcomma}\ Ycoord\ {\isacharequal}\ {\isadigit{9}}{\isadigit{7}}{\isasymrparr}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}pt{\isadigit{2}}\ {\isasymequiv}\ {\isasymlparr}Xcoord\ {\isacharequal}\ {\isacharminus}{\isadigit{4}}{\isadigit{5}}{\isacharcomma}\ Ycoord\ {\isacharequal}\ {\isadigit{9}}{\isadigit{7}}{\isasymrparr}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% For each field, there is a \emph{selector}\index{selector!record} function of the same name. For example, if \isa{p} has type \isa{point} then \isa{Xcoord\ p} denotes the value of the \isa{Xcoord} field of~\isa{p}. Expressions involving field selection of explicit records are simplified automatically:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}Xcoord\ {\isasymlparr}Xcoord\ {\isacharequal}\ a{\isacharcomma}\ Ycoord\ {\isacharequal}\ b{\isasymrparr}\ {\isacharequal}\ a{\isachardoublequote}\isanewline -\ \ \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isamarkupfalse% +\isacommand{by}\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% The \emph{update}\index{update!record} operation is functional. For example, \isa{p{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ {\isadigit{0}}{\isasymrparr}} is a record whose \isa{Xcoord} value is zero and whose \isa{Ycoord} value is copied from~\isa{p}. Updates of explicit records are also simplified automatically:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isasymlparr}Xcoord\ {\isacharequal}\ a{\isacharcomma}\ Ycoord\ {\isacharequal}\ b{\isasymrparr}{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ {\isadigit{0}}{\isasymrparr}\ {\isacharequal}\isanewline \ \ \ \ {\isasymlparr}Xcoord\ {\isacharequal}\ {\isadigit{0}}{\isacharcomma}\ Ycoord\ {\isacharequal}\ b{\isasymrparr}{\isachardoublequote}\isanewline -\ \ \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isamarkupfalse% +\isacommand{by}\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \begin{warn} @@ -109,21 +148,21 @@ Now, let us define coloured points (type \isa{cpoint}) to be points extended with a field \isa{col} of type \isa{colour}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{datatype}\ colour\ {\isacharequal}\ Red\ {\isacharbar}\ Green\ {\isacharbar}\ Blue\isanewline \isanewline \isamarkupfalse% \isacommand{record}\ cpoint\ {\isacharequal}\ point\ {\isacharplus}\isanewline -\ \ col\ {\isacharcolon}{\isacharcolon}\ colour\isamarkupfalse% +\ \ col\ {\isacharcolon}{\isacharcolon}\ colour\isamarkuptrue% % \begin{isamarkuptext}% The fields of this new type are \isa{Xcoord}, \isa{Ycoord} and \isa{col}, in that order.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{constdefs}\isanewline \ \ cpt{\isadigit{1}}\ {\isacharcolon}{\isacharcolon}\ cpoint\isanewline -\ \ {\isachardoublequote}cpt{\isadigit{1}}\ {\isasymequiv}\ {\isasymlparr}Xcoord\ {\isacharequal}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharcomma}\ Ycoord\ {\isacharequal}\ {\isadigit{2}}{\isadigit{3}}{\isacharcomma}\ col\ {\isacharequal}\ Green{\isasymrparr}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}cpt{\isadigit{1}}\ {\isasymequiv}\ {\isasymlparr}Xcoord\ {\isacharequal}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isacharcomma}\ Ycoord\ {\isacharequal}\ {\isadigit{2}}{\isadigit{3}}{\isacharcomma}\ col\ {\isacharequal}\ Green{\isasymrparr}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \medskip We can define generic operations that work on arbitrary @@ -136,10 +175,23 @@ \isa{unit}. Within the record brackets, you can refer to the \isa{more} field by writing ``\isa{{\isasymdots}}'' (three dots):% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}Xcoord\ {\isasymlparr}Xcoord\ {\isacharequal}\ a{\isacharcomma}\ Ycoord\ {\isacharequal}\ b{\isacharcomma}\ {\isasymdots}\ {\isacharequal}\ p{\isasymrparr}\ {\isacharequal}\ a{\isachardoublequote}\isanewline -\ \ \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isamarkupfalse% +\isacommand{by}\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% This lemma applies to any record whose first two fields are \isa{Xcoord} and~\isa{Ycoord}. Note that \isa{{\isasymlparr}Xcoord\ {\isacharequal}\ a{\isacharcomma}\ Ycoord\ {\isacharequal}\ b{\isacharcomma}\ {\isasymdots}\ {\isacharequal}\ {\isacharparenleft}{\isacharparenright}{\isasymrparr}} is exactly the same as \isa{{\isasymlparr}Xcoord\ {\isacharequal}\ a{\isacharcomma}\ Ycoord\ {\isacharequal}\ b{\isasymrparr}}. Selectors and updates are always polymorphic wrt.\ the @@ -149,10 +201,23 @@ The \isa{more} pseudo-field may be manipulated directly as well, but the identifier needs to be qualified:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}point{\isachardot}more\ cpt{\isadigit{1}}\ {\isacharequal}\ {\isasymlparr}col\ {\isacharequal}\ Green{\isasymrparr}{\isachardoublequote}\isanewline -\ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ add{\isacharcolon}\ cpt{\isadigit{1}}{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}simp\ add{\isacharcolon}\ cpt{\isadigit{1}}{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% We see that the colour part attached to this \isa{point} is a @@ -178,12 +243,12 @@ In the following example we define two operations --- methods, if we regard records as objects --- to get and set any point's \isa{Xcoord} field.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{constdefs}\isanewline \ \ getX\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ point{\isacharunderscore}scheme\ {\isasymRightarrow}\ int{\isachardoublequote}\isanewline \ \ {\isachardoublequote}getX\ r\ {\isasymequiv}\ Xcoord\ r{\isachardoublequote}\isanewline \ \ setX\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ point{\isacharunderscore}scheme\ {\isasymRightarrow}\ int\ {\isasymRightarrow}\ {\isacharprime}a\ point{\isacharunderscore}scheme{\isachardoublequote}\isanewline -\ \ {\isachardoublequote}setX\ r\ a\ {\isasymequiv}\ r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isasymrparr}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}setX\ r\ a\ {\isasymequiv}\ r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isasymrparr}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% Here is a generic method that modifies a point, incrementing its @@ -191,20 +256,33 @@ are copied across. It works for any record type scheme derived from \isa{point} (including \isa{cpoint} etc.):% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{constdefs}\isanewline \ \ incX\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharprime}a\ point{\isacharunderscore}scheme\ {\isasymRightarrow}\ {\isacharprime}a\ point{\isacharunderscore}scheme{\isachardoublequote}\isanewline \ \ {\isachardoublequote}incX\ r\ {\isasymequiv}\isanewline -\ \ \ \ {\isasymlparr}Xcoord\ {\isacharequal}\ Xcoord\ r\ {\isacharplus}\ {\isadigit{1}}{\isacharcomma}\ Ycoord\ {\isacharequal}\ Ycoord\ r{\isacharcomma}\ {\isasymdots}\ {\isacharequal}\ point{\isachardot}more\ r{\isasymrparr}{\isachardoublequote}\isamarkupfalse% +\ \ \ \ {\isasymlparr}Xcoord\ {\isacharequal}\ Xcoord\ r\ {\isacharplus}\ {\isadigit{1}}{\isacharcomma}\ Ycoord\ {\isacharequal}\ Ycoord\ r{\isacharcomma}\ {\isasymdots}\ {\isacharequal}\ point{\isachardot}more\ r{\isasymrparr}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% Generic theorems can be proved about generic methods. This trivial lemma relates \isa{incX} to \isa{getX} and \isa{setX}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}incX\ r\ {\isacharequal}\ setX\ r\ {\isacharparenleft}getX\ r\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequote}\isanewline -\ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ add{\isacharcolon}\ getX{\isacharunderscore}def\ setX{\isacharunderscore}def\ incX{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}simp\ add{\isacharcolon}\ getX{\isacharunderscore}def\ setX{\isacharunderscore}def\ incX{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \begin{warn} @@ -227,20 +305,46 @@ corresponding fields are equal. Concrete record equalities are simplified automatically:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}{\isacharparenleft}{\isasymlparr}Xcoord\ {\isacharequal}\ a{\isacharcomma}\ Ycoord\ {\isacharequal}\ b{\isasymrparr}\ {\isacharequal}\ {\isasymlparr}Xcoord\ {\isacharequal}\ a{\isacharprime}{\isacharcomma}\ Ycoord\ {\isacharequal}\ b{\isacharprime}{\isasymrparr}{\isacharparenright}\ {\isacharequal}\isanewline \ \ \ \ {\isacharparenleft}a\ {\isacharequal}\ a{\isacharprime}\ {\isasymand}\ b\ {\isacharequal}\ b{\isacharprime}{\isacharparenright}{\isachardoublequote}\isanewline -\ \ \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isamarkupfalse% +\isacommand{by}\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% The following equality is similar, but generic, in that \isa{r} can be any instance of \isa{{\isacharprime}a\ point{\isacharunderscore}scheme}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isacharcomma}\ Ycoord\ {\isacharcolon}{\isacharequal}\ b{\isasymrparr}\ {\isacharequal}\ r{\isasymlparr}Ycoord\ {\isacharcolon}{\isacharequal}\ b{\isacharcomma}\ Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isasymrparr}{\isachardoublequote}\isanewline -\ \ \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isamarkupfalse% +\isacommand{by}\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% We see above the syntax for iterated updates. We could equivalently @@ -250,31 +354,70 @@ \index{extensionality!for records} a record is determined entirely by the values of its fields.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}r\ {\isacharequal}\ {\isasymlparr}Xcoord\ {\isacharequal}\ Xcoord\ r{\isacharcomma}\ Ycoord\ {\isacharequal}\ Ycoord\ r{\isasymrparr}{\isachardoublequote}\isanewline -\ \ \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isamarkupfalse% +\isacommand{by}\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% The generic version of this equality includes the pseudo-field \isa{more}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}r\ {\isacharequal}\ {\isasymlparr}Xcoord\ {\isacharequal}\ Xcoord\ r{\isacharcomma}\ Ycoord\ {\isacharequal}\ Ycoord\ r{\isacharcomma}\ {\isasymdots}\ {\isacharequal}\ point{\isachardot}more\ r{\isasymrparr}{\isachardoublequote}\isanewline -\ \ \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isamarkupfalse% +\isacommand{by}\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \medskip The simplifier can prove many record equalities automatically, but general equality reasoning can be tricky. Consider proving this obvious fact:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isasymrparr}\ {\isacharequal}\ r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isacharprime}{\isasymrparr}\ {\isasymLongrightarrow}\ a\ {\isacharequal}\ a{\isacharprime}{\isachardoublequote}\isanewline -\ \ \isamarkupfalse% +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isamarkupfalse% \isacommand{apply}\ simp{\isacharquery}\isanewline \ \ \isamarkupfalse% -\isacommand{oops}\isamarkupfalse% +\isacommand{oops}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% Here the simplifier can do nothing, since general record equality is @@ -282,10 +425,16 @@ forward step that applies the selector \isa{Xcoord} to both sides of the assumed record equality:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isasymrparr}\ {\isacharequal}\ r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isacharprime}{\isasymrparr}\ {\isasymLongrightarrow}\ a\ {\isacharequal}\ a{\isacharprime}{\isachardoublequote}\isanewline -\ \ \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}drule{\isacharunderscore}tac\ f\ {\isacharequal}\ Xcoord\ \isakeyword{in}\ arg{\isacharunderscore}cong{\isacharparenright}\isamarkupfalse% +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isamarkupfalse% +\isacommand{apply}\ {\isacharparenleft}drule{\isacharunderscore}tac\ f\ {\isacharequal}\ Xcoord\ \isakeyword{in}\ arg{\isacharunderscore}cong{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -294,17 +443,30 @@ Now, \isa{simp} will reduce the assumption to the desired conclusion.% \end{isamarkuptxt}% -\ \ \isamarkuptrue% +\ \ \isamarkupfalse% \isacommand{apply}\ simp\isanewline \ \ \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% The \isa{cases} method is preferable to such a forward proof. We state the desired lemma again:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ {\isachardoublequote}r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isasymrparr}\ {\isacharequal}\ r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isacharprime}{\isasymrparr}\ {\isasymLongrightarrow}\ a\ {\isacharequal}\ a{\isacharprime}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{lemma}\ {\isachardoublequote}r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isasymrparr}\ {\isacharequal}\ r{\isasymlparr}Xcoord\ {\isacharcolon}{\isacharequal}\ a{\isacharprime}{\isasymrparr}\ {\isasymLongrightarrow}\ a\ {\isacharequal}\ a{\isacharprime}{\isachardoublequote}\isamarkupfalse% % \begin{isamarkuptxt}% The \methdx{cases} method adds an equality to replace the @@ -312,8 +474,8 @@ fields. It even includes the pseudo-field \isa{more}, since the record equality stated here is generic for all extensions.% \end{isamarkuptxt}% -\ \ \isamarkuptrue% -\isacommand{apply}\ {\isacharparenleft}cases\ r{\isacharparenright}\isamarkupfalse% +\ \ \isamarkupfalse% +\isacommand{apply}\ {\isacharparenleft}cases\ r{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -326,10 +488,17 @@ record equality can be replaced by equality of the corresponding fields (due to injectivity).% \end{isamarkuptxt}% -\ \ \isamarkuptrue% +\ \ \isamarkupfalse% \isacommand{apply}\ simp\isanewline \ \ \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% The generic cases method does not admit references to locally bound @@ -402,10 +571,10 @@ \isa{pt{\isadigit{1}}} with a colour value, which is converted into an appropriate record fragment by \isa{cpoint{\isachardot}fields}.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{constdefs}\isanewline \ \ cpt{\isadigit{2}}\ {\isacharcolon}{\isacharcolon}\ cpoint\isanewline -\ \ {\isachardoublequote}cpt{\isadigit{2}}\ {\isasymequiv}\ point{\isachardot}extend\ pt{\isadigit{1}}\ {\isacharparenleft}cpoint{\isachardot}fields\ Green{\isacharparenright}{\isachardoublequote}\isamarkupfalse% +\ \ {\isachardoublequote}cpt{\isadigit{2}}\ {\isasymequiv}\ point{\isachardot}extend\ pt{\isadigit{1}}\ {\isacharparenleft}cpoint{\isachardot}fields\ Green{\isacharparenright}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% The coloured points \isa{cpt{\isadigit{1}}} and \isa{cpt{\isadigit{2}}} are equal. The @@ -413,29 +582,55 @@ omit the definition of~\isa{pt{\isadigit{1}}} in order to reveal the underlying comparison on type \isa{point}.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}cpt{\isadigit{1}}\ {\isacharequal}\ cpt{\isadigit{2}}{\isachardoublequote}\isanewline -\ \ \isamarkupfalse% -\isacommand{apply}\ {\isacharparenleft}simp\ add{\isacharcolon}\ cpt{\isadigit{1}}{\isacharunderscore}def\ cpt{\isadigit{2}}{\isacharunderscore}def\ point{\isachardot}defs\ cpoint{\isachardot}defs{\isacharparenright}\isamarkupfalse% +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isamarkupfalse% +\isacommand{apply}\ {\isacharparenleft}simp\ add{\isacharcolon}\ cpt{\isadigit{1}}{\isacharunderscore}def\ cpt{\isadigit{2}}{\isacharunderscore}def\ point{\isachardot}defs\ cpoint{\isachardot}defs{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% \ {\isadigit{1}}{\isachardot}\ Xcoord\ pt{\isadigit{1}}\ {\isacharequal}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}\ {\isasymand}\ Ycoord\ pt{\isadigit{1}}\ {\isacharequal}\ {\isadigit{2}}{\isadigit{3}}% \end{isabelle}% \end{isamarkuptxt}% -\ \ \isamarkuptrue% +\ \ \isamarkupfalse% \isacommand{apply}\ {\isacharparenleft}simp\ add{\isacharcolon}\ pt{\isadigit{1}}{\isacharunderscore}def{\isacharparenright}\isanewline \ \ \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% In the example below, a coloured point is truncated to leave a point. We use the \isa{truncate} function of the target record.% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}point{\isachardot}truncate\ cpt{\isadigit{2}}\ {\isacharequal}\ pt{\isadigit{1}}{\isachardoublequote}\isanewline -\ \ \isamarkupfalse% -\isacommand{by}\ {\isacharparenleft}simp\ add{\isacharcolon}\ pt{\isadigit{1}}{\isacharunderscore}def\ cpt{\isadigit{2}}{\isacharunderscore}def\ point{\isachardot}defs{\isacharparenright}\isamarkupfalse% +% +\isadelimproof +\ \ % +\endisadelimproof +% +\isatagproof +\isamarkupfalse% +\isacommand{by}\ {\isacharparenleft}simp\ add{\isacharcolon}\ pt{\isadigit{1}}{\isacharunderscore}def\ cpt{\isadigit{2}}{\isacharunderscore}def\ point{\isachardot}defs{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \begin{exercise} @@ -451,8 +646,19 @@ \end{exercise} \index{records|)}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/Types/document/Typedefs.tex --- a/doc-src/TutorialI/Types/document/Typedefs.tex Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/Types/document/Typedefs.tex Tue Aug 16 13:42:23 2005 +0200 @@ -1,7 +1,20 @@ % \begin{isabellebody}% \def\isabellecontext{Typedefs}% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isamarkuptrue% % \isamarkupsection{Introducing New Types% } @@ -32,8 +45,8 @@ The most trivial way of introducing a new type is by a \textbf{type declaration}:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{typedecl}\ my{\isacharunderscore}new{\isacharunderscore}type\isamarkupfalse% +\isamarkupfalse% +\isacommand{typedecl}\ my{\isacharunderscore}new{\isacharunderscore}type\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -52,9 +65,9 @@ together with its properties: declare the type and state its properties as axioms. Example:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{axioms}\isanewline -just{\isacharunderscore}one{\isacharcolon}\ {\isachardoublequote}{\isasymexists}x{\isacharcolon}{\isacharcolon}my{\isacharunderscore}new{\isacharunderscore}type{\isachardot}\ {\isasymforall}y{\isachardot}\ x\ {\isacharequal}\ y{\isachardoublequote}\isamarkupfalse% +just{\isacharunderscore}one{\isacharcolon}\ {\isachardoublequote}{\isasymexists}x{\isacharcolon}{\isacharcolon}my{\isacharunderscore}new{\isacharunderscore}type{\isachardot}\ {\isasymforall}y{\isachardot}\ x\ {\isacharequal}\ y{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -85,8 +98,14 @@ Let us work a simple example, the definition of a three-element type. It is easily represented by the first three natural numbers:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{typedef}\ three\ {\isacharequal}\ {\isachardoublequote}{\isacharbraceleft}{\isadigit{0}}{\isacharcolon}{\isacharcolon}nat{\isacharcomma}\ {\isadigit{1}}{\isacharcomma}\ {\isadigit{2}}{\isacharbraceright}{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{typedef}\ three\ {\isacharequal}\ {\isachardoublequote}{\isacharbraceleft}{\isadigit{0}}{\isacharcolon}{\isacharcolon}nat{\isacharcomma}\ {\isadigit{1}}{\isacharcomma}\ {\isadigit{2}}{\isacharbraceright}{\isachardoublequote}\isamarkupfalse% % \begin{isamarkuptxt}% \noindent @@ -98,10 +117,17 @@ Fortunately, this is easy enough to show, even \isa{auto} could do it. In general, one has to provide a witness, in our case 0:% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}rule{\isacharunderscore}tac\ x\ {\isacharequal}\ {\isadigit{0}}\ \isakeyword{in}\ exI{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{by}\ simp\isamarkupfalse% +\isacommand{by}\ simp% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% This type definition introduces the new type \isa{three} and asserts @@ -151,14 +177,14 @@ In our example it suffices to give the three elements of type \isa{three} names:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{constdefs}\isanewline \ \ A{\isacharcolon}{\isacharcolon}\ three\isanewline \ {\isachardoublequote}A\ {\isasymequiv}\ Abs{\isacharunderscore}three\ {\isadigit{0}}{\isachardoublequote}\isanewline \ \ B{\isacharcolon}{\isacharcolon}\ three\isanewline \ {\isachardoublequote}B\ {\isasymequiv}\ Abs{\isacharunderscore}three\ {\isadigit{1}}{\isachardoublequote}\isanewline \ \ C\ {\isacharcolon}{\isacharcolon}\ three\isanewline -\ {\isachardoublequote}C\ {\isasymequiv}\ Abs{\isacharunderscore}three\ {\isadigit{2}}{\isachardoublequote}\isamarkupfalse% +\ {\isachardoublequote}C\ {\isasymequiv}\ Abs{\isacharunderscore}three\ {\isadigit{2}}{\isachardoublequote}\isamarkuptrue% % \begin{isamarkuptext}% So far, everything was easy. But it is clear that reasoning about \isa{three} will be hell if we have to go back to \isa{nat} every time. Thus our @@ -197,10 +223,23 @@ if we expand their definitions and rewrite with the injectivity of \isa{Abs{\isacharunderscore}three}:% \end{isamarkuptext}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{lemma}\ {\isachardoublequote}A\ {\isasymnoteq}\ B\ {\isasymand}\ B\ {\isasymnoteq}\ A\ {\isasymand}\ A\ {\isasymnoteq}\ C\ {\isasymand}\ C\ {\isasymnoteq}\ A\ {\isasymand}\ B\ {\isasymnoteq}\ C\ {\isasymand}\ C\ {\isasymnoteq}\ B{\isachardoublequote}\isanewline +% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkupfalse% -\isacommand{by}{\isacharparenleft}simp\ add{\isacharcolon}\ Abs{\isacharunderscore}three{\isacharunderscore}inject\ A{\isacharunderscore}def\ B{\isacharunderscore}def\ C{\isacharunderscore}def\ three{\isacharunderscore}def{\isacharparenright}\isamarkupfalse% +\isacommand{by}{\isacharparenleft}simp\ add{\isacharcolon}\ Abs{\isacharunderscore}three{\isacharunderscore}inject\ A{\isacharunderscore}def\ B{\isacharunderscore}def\ C{\isacharunderscore}def\ three{\isacharunderscore}def{\isacharparenright}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -211,14 +250,20 @@ (where \isa{x} is of type \isa{three}) it suffices to prove \isa{P\ A}, \isa{P\ B} and \isa{P\ C}:% \end{isamarkuptext}% +\isamarkupfalse% +\isacommand{lemma}\ three{\isacharunderscore}cases{\isacharcolon}\ {\isachardoublequote}{\isasymlbrakk}\ P\ A{\isacharsemicolon}\ P\ B{\isacharsemicolon}\ P\ C\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ x{\isachardoublequote}% +\isadelimproof +% +\endisadelimproof +% +\isatagproof \isamarkuptrue% -\isacommand{lemma}\ three{\isacharunderscore}cases{\isacharcolon}\ {\isachardoublequote}{\isasymlbrakk}\ P\ A{\isacharsemicolon}\ P\ B{\isacharsemicolon}\ P\ C\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ x{\isachardoublequote}\isamarkupfalse% % \begin{isamarkuptxt}% \noindent Again this follows easily from a pre-proved general theorem:% \end{isamarkuptxt}% -\isamarkuptrue% -\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ x\ rule{\isacharcolon}\ Abs{\isacharunderscore}three{\isacharunderscore}induct{\isacharparenright}\isamarkupfalse% +\isamarkupfalse% +\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ x\ rule{\isacharcolon}\ Abs{\isacharunderscore}three{\isacharunderscore}induct{\isacharparenright}\isamarkuptrue% % \begin{isamarkuptxt}% \begin{isabelle}% @@ -227,10 +272,17 @@ Simplification with \isa{three{\isacharunderscore}def} leads to the disjunction \isa{y\ {\isacharequal}\ {\isadigit{0}}\ {\isasymor}\ y\ {\isacharequal}\ {\isadigit{1}}\ {\isasymor}\ y\ {\isacharequal}\ {\isadigit{2}}} which \isa{auto} separates into three subgoals, each of which is easily solved by simplification:% \end{isamarkuptxt}% -\isamarkuptrue% +\isamarkupfalse% \isacommand{apply}{\isacharparenleft}auto\ simp\ add{\isacharcolon}\ three{\isacharunderscore}def\ A{\isacharunderscore}def\ B{\isacharunderscore}def\ C{\isacharunderscore}def{\isacharparenright}\isanewline \isamarkupfalse% -\isacommand{done}\isamarkupfalse% +\isacommand{done}% +\endisatagproof +{\isafoldproof}% +% +\isadelimproof +% +\endisadelimproof +\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -240,8 +292,8 @@ The attentive reader has realized long ago that the above lengthy definition can be collapsed into one line:% \end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\ better{\isacharunderscore}three\ {\isacharequal}\ A\ {\isacharbar}\ B\ {\isacharbar}\ C\isamarkupfalse% +\isamarkupfalse% +\isacommand{datatype}\ better{\isacharunderscore}three\ {\isacharequal}\ A\ {\isacharbar}\ B\ {\isacharbar}\ C\isamarkuptrue% % \begin{isamarkuptext}% \noindent @@ -270,8 +322,19 @@ \index{typedecl@\isacommand {typedef} (command)|)}% \index{types!defining|)}% \end{isamarkuptext}% -\isamarkuptrue% -\isamarkupfalse% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r eacce1cd716a -r 05fc32a23b8b doc-src/TutorialI/isabelle.sty --- a/doc-src/TutorialI/isabelle.sty Tue Aug 16 13:42:21 2005 +0200 +++ b/doc-src/TutorialI/isabelle.sty Tue Aug 16 13:42:23 2005 +0200 @@ -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}}{}