# HG changeset patch # User nipkow # Date 970565214 -7200 # Node ID e187dacd248f04ff02e481d13ebf48f925b2a5a1 # Parent 8e9a8ede2f112b4b82a29c47a802b71b4da78993 *** empty log message *** diff -r 8e9a8ede2f11 -r e187dacd248f doc-src/TutorialI/CTL/Base.thy --- a/doc-src/TutorialI/CTL/Base.thy Tue Oct 03 01:15:11 2000 +0200 +++ b/doc-src/TutorialI/CTL/Base.thy Tue Oct 03 11:26:54 2000 +0200 @@ -16,35 +16,70 @@ very efficient one. The main proof obligation is to show that the semantics and the model checker agree. -Our models are transition systems. +\underscoreon -START with simple example: mutex or something. +Our models are \emph{transition systems}, i.e.\ sets of \emph{states} with +transitions between them, as shown in this simple example: +\begin{center} +\unitlength.5mm +\thicklines +\begin{picture}(100,60) +\put(50,50){\circle{20}} +\put(50,50){\makebox(0,0){$p,q$}} +\put(61,55){\makebox(0,0)[l]{$s_0$}} +\put(44,42){\vector(-1,-1){26}} +\put(16,18){\vector(1,1){26}} +\put(57,43){\vector(1,-1){26}} +\put(10,10){\circle{20}} +\put(10,10){\makebox(0,0){$q,r$}} +\put(-1,15){\makebox(0,0)[r]{$s_1$}} +\put(20,10){\vector(1,0){60}} +\put(90,10){\circle{20}} +\put(90,10){\makebox(0,0){$r$}} +\put(98, 5){\line(1,0){10}} +\put(108, 5){\line(0,1){10}} +\put(108,15){\vector(-1,0){10}} +\put(91,21){\makebox(0,0)[bl]{$s_2$}} +\end{picture} +\end{center} +Each state has a unique name or number ($s_0,s_1,s_2$), and in each +state certain \emph{atomic propositions} ($p,q,r$) are true. +The aim of temporal logic is to formalize statements such as ``there is no +transition sequence starting from $s_2$ leading to a state where $p$ or $q$ +are true''. Abstracting from this concrete example, we assume there is some type of -atomic propositions +states *} -typedecl atom +typedecl state text{*\noindent -which we merely declare rather than define because it is an implicit parameter of our model. -Of course it would have been more generic to make @{typ atom} a type parameter of everything -but fixing @{typ atom} as above reduces clutter. - -Instead of introducing both a separate type of states and a function -telling us which atoms are true in each state we simply identify each -state with that set of atoms: -*} - -types state = "atom set"; - -text{*\noindent -Finally we declare an arbitrary but fixed transition system, i.e.\ relation between states: +which we merely declare rather than define because it is an implicit +parameter of our model. Of course it would have been more generic to make +@{typ state} a type parameter of everything but fixing @{typ state} as above +reduces clutter. +Similarly we declare an arbitrary but fixed transition system, i.e.\ +relation between states: *} consts M :: "(state \ state)set"; text{*\noindent Again, we could have made @{term M} a parameter of everything. +Finally we introduce a type of atomic propositions *} + +typedecl atom + +text{*\noindent +and a \emph{labelling function} +*} + +consts L :: "state \ atom set" + +text{*\noindent +telling us which atomic propositions are true in each state. +*} + (*<*)end(*>*) diff -r 8e9a8ede2f11 -r e187dacd248f doc-src/TutorialI/CTL/CTL.thy --- a/doc-src/TutorialI/CTL/CTL.thy Tue Oct 03 01:15:11 2000 +0200 +++ b/doc-src/TutorialI/CTL/CTL.thy Tue Oct 03 11:26:54 2000 +0200 @@ -1,7 +1,6 @@ -theory CTL = Main: +(*<*)theory CTL = Base:(*>*) -typedecl atom; -types state = "atom set"; +subsection{*Computation tree logic---CTL*} datatype ctl_form = Atom atom | NOT ctl_form @@ -11,13 +10,12 @@ | AF ctl_form; consts valid :: "state \ ctl_form \ bool" ("(_ \ _)" [80,80] 80) - M :: "(state \ state)set"; constdefs Paths :: "state \ (nat \ state)set" "Paths s \ {p. s = p 0 \ (\i. (p i, p(i+1)) \ M)}"; primrec -"s \ Atom a = (a\s)" +"s \ Atom a = (a \ L s)" "s \ NOT f = (~(s \ f))" "s \ And f g = (s \ f \ s \ g)" "s \ AX f = (\t. (s,t) \ M \ t \ f)" @@ -32,7 +30,7 @@ consts mc :: "ctl_form \ state set"; primrec -"mc(Atom a) = {s. a\s}" +"mc(Atom a) = {s. a \ L s}" "mc(NOT f) = -mc f" "mc(And f g) = mc f \ mc g" "mc(AX f) = {s. \t. (s,t) \ M \ t \ mc f}" @@ -231,4 +229,4 @@ apply(induct_tac f); by(auto simp add: lfp_conv_EF equalityI[OF lfp_subset_AF AF_subset_lfp]); -end; +(*<*)end(*>*) diff -r 8e9a8ede2f11 -r e187dacd248f doc-src/TutorialI/CTL/PDL.thy --- a/doc-src/TutorialI/CTL/PDL.thy Tue Oct 03 01:15:11 2000 +0200 +++ b/doc-src/TutorialI/CTL/PDL.thy Tue Oct 03 11:26:54 2000 +0200 @@ -2,36 +2,61 @@ subsection{*Propositional dynmic logic---PDL*} -datatype ctl_form = Atom atom - | NOT ctl_form - | And ctl_form ctl_form - | AX ctl_form - | EF ctl_form; +text{* +The formulae of PDL are built up from atomic propositions via the customary +propositional connectives of negation and conjunction and the two temporal +connectives @{text AX} and @{text EF}. Since formulae are essentially +(syntax) trees, they are naturally modelled as a datatype: +*} -consts valid :: "state \ ctl_form \ bool" ("(_ \ _)" [80,80] 80) - M :: "(state \ state)set"; +datatype pdl_form = Atom atom + | NOT pdl_form + | And pdl_form pdl_form + | AX pdl_form + | EF pdl_form; + +text{*\noindent +The meaning of these formulae is given by saying which formula is true in +which state: +*} + +consts valid :: "state \ pdl_form \ bool" ("(_ \ _)" [80,80] 80) primrec -"s \ Atom a = (a\s)" +"s \ Atom a = (a \ L s)" "s \ NOT f = (\(s \ f))" "s \ And f g = (s \ f \ s \ g)" "s \ AX f = (\t. (s,t) \ M \ t \ f)" "s \ EF f = (\t. (s,t) \ M^* \ t \ f)"; -consts mc :: "ctl_form \ state set"; +text{* +Now we come to the model checker itself. It maps a formula into the set of +states where the formula is true and is defined by recursion over the syntax: +*} + +consts mc :: "pdl_form \ state set"; primrec -"mc(Atom a) = {s. a\s}" +"mc(Atom a) = {s. a \ L s}" "mc(NOT f) = -mc f" -"mc(And f g) = mc f Int mc g" +"mc(And f g) = mc f \ mc g" "mc(AX f) = {s. \t. (s,t) \ M \ t \ mc f}" -"mc(EF f) = lfp(\T. mc f \ {s. \t. (s,t)\M \ t\T})"; +"mc(EF f) = lfp(\T. mc f \ M^-1 ^^ T)" + -lemma mono_lemma: "mono(\T. A \ {s. \t. (s,t)\M \ t\T})"; +text{* +Only the equation for @{term EF} deserves a comment: the postfix @{text"^-1"} +and the infix @{text"^^"} are predefined and denote the converse of a +relation and the application of a relation to a set. Thus @{term "M^-1 ^^ T"} +is the set of all predecessors of @{term T} and @{term"mc(EF f)"} is the least +set @{term T} containing @{term"mc f"} and all predecessors of @{term T}. +*} + +lemma mono_lemma: "mono(\T. A \ M^-1 ^^ T)" apply(rule monoI); by(blast); lemma lfp_conv_EF: -"lfp(\T. A \ {s. \t. (s,t)\M \ t\T}) = {s. \t. (s,t) \ M^* \ t \ A}"; +"lfp(\T. A \ M^-1 ^^ T) = {s. \t. (s,t) \ M^* \ t \ A}"; apply(rule equalityI); apply(rule subsetI); apply(simp); @@ -53,5 +78,4 @@ theorem "mc f = {s. s \ f}"; apply(induct_tac f); by(auto simp add:lfp_conv_EF); - -end; +(*<*)end(*>*) diff -r 8e9a8ede2f11 -r e187dacd248f doc-src/TutorialI/CTL/ctl.tex --- a/doc-src/TutorialI/CTL/ctl.tex Tue Oct 03 01:15:11 2000 +0200 +++ b/doc-src/TutorialI/CTL/ctl.tex Tue Oct 03 11:26:54 2000 +0200 @@ -1,3 +1,3 @@ \input{CTL/document/Base.tex} \input{CTL/document/PDL.tex} - +\input{CTL/document/CTL.tex} diff -r 8e9a8ede2f11 -r e187dacd248f doc-src/TutorialI/CTL/document/Base.tex --- a/doc-src/TutorialI/CTL/document/Base.tex Tue Oct 03 01:15:11 2000 +0200 +++ b/doc-src/TutorialI/CTL/document/Base.tex Tue Oct 03 11:26:54 2000 +0200 @@ -18,33 +18,66 @@ very efficient one. The main proof obligation is to show that the semantics and the model checker agree. -Our models are transition systems. +\underscoreon -START with simple example: mutex or something. +Our models are \emph{transition systems}, i.e.\ sets of \emph{states} with +transitions between them, as shown in this simple example: +\begin{center} +\unitlength.5mm +\thicklines +\begin{picture}(100,60) +\put(50,50){\circle{20}} +\put(50,50){\makebox(0,0){$p,q$}} +\put(61,55){\makebox(0,0)[l]{$s_0$}} +\put(44,42){\vector(-1,-1){26}} +\put(16,18){\vector(1,1){26}} +\put(57,43){\vector(1,-1){26}} +\put(10,10){\circle{20}} +\put(10,10){\makebox(0,0){$q,r$}} +\put(-1,15){\makebox(0,0)[r]{$s_1$}} +\put(20,10){\vector(1,0){60}} +\put(90,10){\circle{20}} +\put(90,10){\makebox(0,0){$r$}} +\put(98, 5){\line(1,0){10}} +\put(108, 5){\line(0,1){10}} +\put(108,15){\vector(-1,0){10}} +\put(91,21){\makebox(0,0)[bl]{$s_2$}} +\end{picture} +\end{center} +Each state has a unique name or number ($s_0,s_1,s_2$), and in each +state certain \emph{atomic propositions} ($p,q,r$) are true. +The aim of temporal logic is to formalize statements such as ``there is no +transition sequence starting from $s_2$ leading to a state where $p$ or $q$ +are true''. Abstracting from this concrete example, we assume there is some type of -atomic propositions% +states% +\end{isamarkuptext}% +\isacommand{typedecl}\ state% +\begin{isamarkuptext}% +\noindent +which we merely declare rather than define because it is an implicit +parameter of our model. Of course it would have been more generic to make +\isa{state} a type parameter of everything but fixing \isa{state} as above +reduces clutter. +Similarly we declare an arbitrary but fixed transition system, i.e.\ +relation between states:% +\end{isamarkuptext}% +\isacommand{consts}\ M\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}state\ {\isasymtimes}\ state{\isacharparenright}set{\isachardoublequote}% +\begin{isamarkuptext}% +\noindent +Again, we could have made \isa{M} a parameter of everything. +Finally we introduce a type of atomic propositions% \end{isamarkuptext}% \isacommand{typedecl}\ atom% \begin{isamarkuptext}% \noindent -which we merely declare rather than define because it is an implicit parameter of our model. -Of course it would have been more generic to make \isa{atom} a type parameter of everything -but fixing \isa{atom} as above reduces clutter. - -Instead of introducing both a separate type of states and a function -telling us which atoms are true in each state we simply identify each -state with that set of atoms:% +and a \emph{labelling function}% \end{isamarkuptext}% -\isacommand{types}\ state\ {\isacharequal}\ {\isachardoublequote}atom\ set{\isachardoublequote}% +\isacommand{consts}\ L\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}state\ {\isasymRightarrow}\ atom\ set{\isachardoublequote}% \begin{isamarkuptext}% \noindent -Finally we declare an arbitrary but fixed transition system, i.e.\ relation between states:% -\end{isamarkuptext}% -\isacommand{consts}\ M\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}state\ {\isasymtimes}\ state{\isacharparenright}set{\isachardoublequote}% -\begin{isamarkuptext}% -\noindent -Again, we could have made \isa{M} a parameter of everything.% +telling us which atomic propositions are true in each state.% \end{isamarkuptext}% \end{isabellebody}% %%% Local Variables: diff -r 8e9a8ede2f11 -r e187dacd248f doc-src/TutorialI/CTL/document/CTL.tex --- a/doc-src/TutorialI/CTL/document/CTL.tex Tue Oct 03 01:15:11 2000 +0200 +++ b/doc-src/TutorialI/CTL/document/CTL.tex Tue Oct 03 11:26:54 2000 +0200 @@ -1,11 +1,8 @@ % \begin{isabellebody}% \def\isabellecontext{CTL}% -\isacommand{theory}\ CTL\ {\isacharequal}\ Main{\isacharcolon}\isanewline -\isanewline -\isacommand{typedecl}\ atom\isanewline -\isacommand{types}\ state\ {\isacharequal}\ {\isachardoublequote}atom\ set{\isachardoublequote}\isanewline -\isanewline +% +\isamarkupsubsection{Computation tree logic---CTL} \isacommand{datatype}\ ctl{\isacharunderscore}form\ {\isacharequal}\ Atom\ atom\isanewline \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ NOT\ ctl{\isacharunderscore}form\isanewline \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ And\ ctl{\isacharunderscore}form\ ctl{\isacharunderscore}form\isanewline @@ -14,13 +11,12 @@ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ AF\ ctl{\isacharunderscore}form\isanewline \isanewline \isacommand{consts}\ valid\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}state\ {\isasymRightarrow}\ ctl{\isacharunderscore}form\ {\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}\isanewline -\ \ \ \ \ \ \ M\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}state\ {\isasymtimes}\ state{\isacharparenright}set{\isachardoublequote}\isanewline \isanewline \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}\isanewline \isanewline \isacommand{primrec}\isanewline -{\isachardoublequote}s\ {\isasymTurnstile}\ Atom\ a\ \ {\isacharequal}\ \ {\isacharparenleft}a{\isasymin}s{\isacharparenright}{\isachardoublequote}\isanewline +{\isachardoublequote}s\ {\isasymTurnstile}\ Atom\ a\ \ {\isacharequal}\ \ {\isacharparenleft}a\ {\isasymin}\ L\ s{\isacharparenright}{\isachardoublequote}\isanewline {\isachardoublequote}s\ {\isasymTurnstile}\ NOT\ f\ \ \ {\isacharequal}\ {\isacharparenleft}{\isachartilde}{\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 @@ -35,7 +31,7 @@ \isanewline \isacommand{consts}\ mc\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}ctl{\isacharunderscore}form\ {\isasymRightarrow}\ state\ set{\isachardoublequote}\isanewline \isacommand{primrec}\isanewline -{\isachardoublequote}mc{\isacharparenleft}Atom\ a{\isacharparenright}\ \ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ a{\isasymin}s{\isacharbraceright}{\isachardoublequote}\isanewline +{\isachardoublequote}mc{\isacharparenleft}Atom\ a{\isacharparenright}\ \ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ a\ {\isasymin}\ L\ s{\isacharbraceright}{\isachardoublequote}\isanewline {\isachardoublequote}mc{\isacharparenleft}NOT\ 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 @@ -138,11 +134,11 @@ \isacommand{apply}{\isacharparenleft}intro\ strip{\isacharparenright}\isanewline \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ i{\isacharparenright}\isanewline \ \isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isanewline -\ \isacommand{apply}{\isacharparenleft}fast\ intro{\isacharcolon}someI\isadigit{2}EX{\isacharparenright}\isanewline +\ \isacommand{apply}{\isacharparenleft}fast\ intro{\isacharcolon}someI\isadigit{2}{\isacharunderscore}ex{\isacharparenright}\isanewline \isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isanewline -\isacommand{apply}{\isacharparenleft}rule\ someI\isadigit{2}EX{\isacharparenright}\isanewline +\isacommand{apply}{\isacharparenleft}rule\ someI\isadigit{2}{\isacharunderscore}ex{\isacharparenright}\isanewline \ \isacommand{apply}{\isacharparenleft}blast{\isacharparenright}\isanewline -\isacommand{apply}{\isacharparenleft}rule\ someI\isadigit{2}EX{\isacharparenright}\isanewline +\isacommand{apply}{\isacharparenleft}rule\ someI\isadigit{2}{\isacharunderscore}ex{\isacharparenright}\isanewline \ \isacommand{apply}{\isacharparenleft}blast{\isacharparenright}\isanewline \isacommand{by}{\isacharparenleft}blast{\isacharparenright}\isanewline \isanewline @@ -158,11 +154,11 @@ \isacommand{apply}{\isacharparenleft}intro\ strip{\isacharparenright}\isanewline \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ i{\isacharparenright}\isanewline \ \isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isanewline -\ \isacommand{apply}{\isacharparenleft}fast\ intro{\isacharcolon}someI\isadigit{2}EX{\isacharparenright}\isanewline +\ \isacommand{apply}{\isacharparenleft}fast\ intro{\isacharcolon}someI\isadigit{2}{\isacharunderscore}ex{\isacharparenright}\isanewline \isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isanewline -\isacommand{apply}{\isacharparenleft}rule\ someI\isadigit{2}EX{\isacharparenright}\isanewline +\isacommand{apply}{\isacharparenleft}rule\ someI\isadigit{2}{\isacharunderscore}ex{\isacharparenright}\isanewline \ \isacommand{apply}{\isacharparenleft}blast{\isacharparenright}\isanewline -\isacommand{apply}{\isacharparenleft}rule\ someI\isadigit{2}EX{\isacharparenright}\isanewline +\isacommand{apply}{\isacharparenleft}rule\ someI\isadigit{2}{\isacharunderscore}ex{\isacharparenright}\isanewline \ \isacommand{apply}{\isacharparenleft}blast{\isacharparenright}\isanewline \isacommand{by}{\isacharparenleft}blast{\isacharparenright}\isanewline \isanewline @@ -219,8 +215,6 @@ \isacommand{theorem}\ {\isachardoublequote}mc\ f\ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ s\ {\isasymTurnstile}\ f{\isacharbraceright}{\isachardoublequote}\isanewline \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ f{\isacharparenright}\isanewline \isacommand{by}{\isacharparenleft}auto\ simp\ add{\isacharcolon}\ lfp{\isacharunderscore}conv{\isacharunderscore}EF\ equalityI{\isacharbrackleft}OF\ lfp{\isacharunderscore}subset{\isacharunderscore}AF\ AF{\isacharunderscore}subset{\isacharunderscore}lfp{\isacharbrackright}{\isacharparenright}\isanewline -\isanewline -\isacommand{end}\isanewline \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r 8e9a8ede2f11 -r e187dacd248f doc-src/TutorialI/CTL/document/PDL.tex --- a/doc-src/TutorialI/CTL/document/PDL.tex Tue Oct 03 01:15:11 2000 +0200 +++ b/doc-src/TutorialI/CTL/document/PDL.tex Tue Oct 03 11:26:54 2000 +0200 @@ -3,36 +3,55 @@ \def\isabellecontext{PDL}% % \isamarkupsubsection{Propositional dynmic logic---PDL} -\isacommand{datatype}\ ctl{\isacharunderscore}form\ {\isacharequal}\ Atom\ atom\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ NOT\ ctl{\isacharunderscore}form\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ And\ ctl{\isacharunderscore}form\ ctl{\isacharunderscore}form\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ AX\ ctl{\isacharunderscore}form\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ EF\ ctl{\isacharunderscore}form\isanewline -\isanewline -\isacommand{consts}\ valid\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}state\ {\isasymRightarrow}\ ctl{\isacharunderscore}form\ {\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}\isanewline -\ \ \ \ \ \ \ M\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}state\ {\isasymtimes}\ state{\isacharparenright}set{\isachardoublequote}\isanewline +% +\begin{isamarkuptext}% +The formulae of PDL are built up from atomic propositions via the customary +propositional connectives of negation and conjunction and the two temporal +connectives \isa{AX} and \isa{EF}. Since formulae are essentially +(syntax) trees, they are naturally modelled as a datatype:% +\end{isamarkuptext}% +\isacommand{datatype}\ pdl{\isacharunderscore}form\ {\isacharequal}\ Atom\ atom\isanewline +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ NOT\ pdl{\isacharunderscore}form\isanewline +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ And\ pdl{\isacharunderscore}form\ pdl{\isacharunderscore}form\isanewline +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ AX\ pdl{\isacharunderscore}form\isanewline +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ EF\ pdl{\isacharunderscore}form% +\begin{isamarkuptext}% +\noindent +The meaning of these formulae is given by saying which formula is true in +which state:% +\end{isamarkuptext}% +\isacommand{consts}\ valid\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}state\ {\isasymRightarrow}\ pdl{\isacharunderscore}form\ {\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}\isanewline \isanewline \isacommand{primrec}\isanewline -{\isachardoublequote}s\ {\isasymTurnstile}\ Atom\ a\ \ {\isacharequal}\ {\isacharparenleft}a{\isasymin}s{\isacharparenright}{\isachardoublequote}\isanewline +{\isachardoublequote}s\ {\isasymTurnstile}\ Atom\ a\ \ {\isacharequal}\ {\isacharparenleft}a\ {\isasymin}\ L\ s{\isacharparenright}{\isachardoublequote}\isanewline {\isachardoublequote}s\ {\isasymTurnstile}\ NOT\ 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{\isacharcircum}{\isacharasterisk}\ {\isasymand}\ t\ {\isasymTurnstile}\ f{\isacharparenright}{\isachardoublequote}\isanewline -\isanewline -\isacommand{consts}\ mc\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}ctl{\isacharunderscore}form\ {\isasymRightarrow}\ state\ set{\isachardoublequote}\isanewline +{\isachardoublequote}s\ {\isasymTurnstile}\ EF\ f\ \ \ \ {\isacharequal}\ {\isacharparenleft}{\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}\ {\isasymand}\ t\ {\isasymTurnstile}\ f{\isacharparenright}{\isachardoublequote}% +\begin{isamarkuptext}% +Now we come to the model checker itself. It maps a formula into the set of +states where the formula is true and is defined by recursion over the syntax:% +\end{isamarkuptext}% +\isacommand{consts}\ mc\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}pdl{\isacharunderscore}form\ {\isasymRightarrow}\ state\ set{\isachardoublequote}\isanewline \isacommand{primrec}\isanewline -{\isachardoublequote}mc{\isacharparenleft}Atom\ a{\isacharparenright}\ \ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ a{\isasymin}s{\isacharbraceright}{\isachardoublequote}\isanewline +{\isachardoublequote}mc{\isacharparenleft}Atom\ a{\isacharparenright}\ \ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ a\ {\isasymin}\ L\ s{\isacharbraceright}{\isachardoublequote}\isanewline {\isachardoublequote}mc{\isacharparenleft}NOT\ f{\isacharparenright}\ \ \ {\isacharequal}\ {\isacharminus}mc\ f{\isachardoublequote}\isanewline -{\isachardoublequote}mc{\isacharparenleft}And\ f\ g{\isacharparenright}\ {\isacharequal}\ mc\ f\ Int\ mc\ g{\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}\ {\isacharbraceleft}s{\isachardot}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}{\isasymin}M\ {\isasymand}\ t{\isasymin}T{\isacharbraceright}{\isacharparenright}{\isachardoublequote}\isanewline -\isanewline -\isacommand{lemma}\ mono{\isacharunderscore}lemma{\isacharcolon}\ {\isachardoublequote}mono{\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ {\isacharbraceleft}s{\isachardot}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}{\isasymin}M\ {\isasymand}\ t{\isasymin}T{\isacharbraceright}{\isacharparenright}{\isachardoublequote}\isanewline +{\isachardoublequote}mc{\isacharparenleft}EF\ f{\isacharparenright}\ \ \ \ {\isacharequal}\ lfp{\isacharparenleft}{\isasymlambda}T{\isachardot}\ mc\ f\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ T{\isacharparenright}{\isachardoublequote}% +\begin{isamarkuptext}% +Only the equation for \isa{EF} deserves a comment: the postfix \isa{{\isacharcircum}{\isacharminus}\isadigit{1}} +and the infix \isa{{\isacharcircum}{\isacharcircum}} are predefined and denote the converse of a +relation and the application of a relation to a set. Thus \isa{M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ T} +is the set of all predecessors of \isa{T} and \isa{mc\ {\isacharparenleft}EF\ f{\isacharparenright}} is the least +set \isa{T} containing \isa{mc\ f} and all predecessors of \isa{T}.% +\end{isamarkuptext}% +\isacommand{lemma}\ mono{\isacharunderscore}lemma{\isacharcolon}\ {\isachardoublequote}mono{\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ T{\isacharparenright}{\isachardoublequote}\isanewline \isacommand{apply}{\isacharparenleft}rule\ monoI{\isacharparenright}\isanewline \isacommand{by}{\isacharparenleft}blast{\isacharparenright}\isanewline \isanewline \isacommand{lemma}\ lfp{\isacharunderscore}conv{\isacharunderscore}EF{\isacharcolon}\isanewline -{\isachardoublequote}lfp{\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ {\isacharbraceleft}s{\isachardot}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}{\isasymin}M\ {\isasymand}\ t{\isasymin}T{\isacharbraceright}{\isacharparenright}\ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A{\isacharbraceright}{\isachardoublequote}\isanewline +{\isachardoublequote}lfp{\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ T{\isacharparenright}\ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A{\isacharbraceright}{\isachardoublequote}\isanewline \isacommand{apply}{\isacharparenleft}rule\ equalityI{\isacharparenright}\isanewline \ \isacommand{apply}{\isacharparenleft}rule\ subsetI{\isacharparenright}\isanewline \ \isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isanewline @@ -53,10 +72,7 @@ \isanewline \isacommand{theorem}\ {\isachardoublequote}mc\ f\ {\isacharequal}\ {\isacharbraceleft}s{\isachardot}\ s\ {\isasymTurnstile}\ f{\isacharbraceright}{\isachardoublequote}\isanewline \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ f{\isacharparenright}\isanewline -\isacommand{by}{\isacharparenleft}auto\ simp\ add{\isacharcolon}lfp{\isacharunderscore}conv{\isacharunderscore}EF{\isacharparenright}\isanewline -\isanewline -\isacommand{end}\isanewline -\end{isabellebody}% +\isacommand{by}{\isacharparenleft}auto\ simp\ add{\isacharcolon}lfp{\isacharunderscore}conv{\isacharunderscore}EF{\isacharparenright}\end{isabellebody}% %%% Local Variables: %%% mode: latex %%% TeX-master: "root" diff -r 8e9a8ede2f11 -r e187dacd248f doc-src/TutorialI/tutorial.tex --- a/doc-src/TutorialI/tutorial.tex Tue Oct 03 01:15:11 2000 +0200 +++ b/doc-src/TutorialI/tutorial.tex Tue Oct 03 11:26:54 2000 +0200 @@ -65,6 +65,7 @@ \input{basics} \input{fp} +%\input{CTL/ctl} \input{Advanced/advanced} %\chapter{The Tricks of the Trade} \input{appendix}