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