# HG changeset patch # User nipkow # Date 970673745 -7200 # Node ID 7cfdf6a330a04a00e685b7835107a3f8adedc294 # Parent 739327964a5c1cf05a6ea7e3eb11f7e69830638b *** empty log message *** diff -r 739327964a5c -r 7cfdf6a330a0 doc-src/TutorialI/CTL/CTL.thy --- a/doc-src/TutorialI/CTL/CTL.thy Tue Oct 03 22:39:49 2000 +0200 +++ b/doc-src/TutorialI/CTL/CTL.thy Wed Oct 04 17:35:45 2000 +0200 @@ -2,64 +2,99 @@ subsection{*Computation tree logic---CTL*} -datatype ctl_form = Atom atom - | NOT ctl_form - | And ctl_form ctl_form - | AX ctl_form - | EF ctl_form - | AF ctl_form; +text{* +The semantics of PDL only needs transitive reflexive closure. +Let us now be a bit more adventurous and introduce a new temporal operator +that goes beyond transitive reflexive closure. We extend the datatype +@{text formula} by a new constructor +*} +(*<*) +datatype formula = Atom atom + | Neg formula + | And formula formula + | AX formula + | EF formula(*>*) + | AF formula -consts valid :: "state \ ctl_form \ bool" ("(_ \ _)" [80,80] 80) +text{*\noindent +which stands for "always in the future": +on all paths, at some point the formula holds. +Introducing the notion of paths (in @{term M}) +*} constdefs Paths :: "state \ (nat \ state)set" -"Paths s \ {p. s = p 0 \ (\i. (p i, p(i+1)) \ M)}"; + "Paths s \ {p. s = p 0 \ (\i. (p i, p(i+1)) \ M)}"; + +text{*\noindent +allows a very succinct definition of the semantics of @{term AF}: +\footnote{Do not be mislead: neither datatypes nor recursive functions can be +extended by new constructors or equations. This is just a trick of the +presentation. In reality one has to define a new datatype and a new function.} +*} +(*<*) +consts valid :: "state \ formula \ bool" ("(_ \ _)" [80,80] 80) primrec "s \ Atom a = (a \ L s)" -"s \ NOT f = (~(s \ f))" +"s \ Neg 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)" +(*>*) "s \ AF f = (\p \ Paths s. \i. p i \ f)"; +text{*\noindent +Model checking @{term AF} involves a function which +is just large enough to warrant a separate definition: +*} + constdefs af :: "state set \ state set \ state set" -"af A T \ A \ {s. \t. (s, t) \ M \ t \ T}"; + "af A T \ A \ {s. \t. (s, t) \ M \ t \ T}"; + +text{*\noindent +This function is monotone in its second argument (and also its first, but +that is irrelevant), and hence @{term"af A"} has a least fixpoint. +*} lemma mono_af: "mono(af A)"; -by(force simp add: af_def intro:monoI); +apply(simp add: mono_def af_def) +by(blast); -consts mc :: "ctl_form \ state set"; +text{*\noindent +Now we can define @{term "mc(AF f)"} as the least set @{term T} that contains +@{term"mc f"} and all states all of whose direct successors are in @{term T}: +*} +(*<*) +consts mc :: "formula \ state set"; primrec "mc(Atom a) = {s. a \ L s}" -"mc(NOT f) = -mc f" +"mc(Neg f) = -mc f" "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)"(*>*) "mc(AF f) = lfp(af(mc f))"; +(*<*) +lemma mono_ef: "mono(\T. A \ M^-1 ^^ T)" +apply(rule monoI) +by(blast) -lemma mono_ef: "mono(\T. A \ {s. \t. (s,t)\M \ t\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}"; +lemma EF_lemma: + "lfp(\T. A \ M^-1 ^^ T) = {s. \t. (s,t) \ M^* \ t \ A}" apply(rule equalityI); apply(rule subsetI); - apply(simp); - apply(erule Lfp.induct); - apply(rule mono_ef); - apply(simp); + apply(simp) + apply(erule Lfp.induct) + apply(rule mono_ef) + apply(simp) apply(blast intro: r_into_rtrancl rtrancl_trans); -apply(rule subsetI); -apply(simp); -apply(erule exE); -apply(erule conjE); -apply(erule_tac P = "t\A" in rev_mp); -apply(erule converse_rtrancl_induct); - apply(rule ssubst [OF lfp_Tarski[OF mono_ef]]); - apply(blast); -apply(rule ssubst [OF lfp_Tarski[OF mono_ef]]); -by(blast); +apply(rule subsetI) +apply(simp, clarify) +apply(erule converse_rtrancl_induct) + apply(rule ssubst[OF lfp_Tarski[OF mono_ef]]) + apply(blast) +apply(rule ssubst[OF lfp_Tarski[OF mono_ef]]) +by(blast) +(*>*) theorem lfp_subset_AF: "lfp(af A) \ {s. \ p \ Paths s. \ i. p i \ A}"; @@ -227,6 +262,6 @@ theorem "mc f = {s. s \ f}"; apply(induct_tac f); -by(auto simp add: lfp_conv_EF equalityI[OF lfp_subset_AF AF_subset_lfp]); +by(auto simp add: EF_lemma equalityI[OF lfp_subset_AF AF_subset_lfp]); (*<*)end(*>*) diff -r 739327964a5c -r 7cfdf6a330a0 doc-src/TutorialI/CTL/PDL.thy --- a/doc-src/TutorialI/CTL/PDL.thy Tue Oct 03 22:39:49 2000 +0200 +++ b/doc-src/TutorialI/CTL/PDL.thy Wed Oct 04 17:35:45 2000 +0200 @@ -9,73 +9,179 @@ (syntax) trees, they are naturally modelled as a datatype: *} -datatype pdl_form = Atom atom - | NOT pdl_form - | And pdl_form pdl_form - | AX pdl_form - | EF pdl_form; +datatype formula = Atom atom + | Neg formula + | And formula formula + | AX formula + | EF formula text{*\noindent +This is almost the same as in the boolean expression case study in +\S\ref{sec:boolex}, except that what used to be called @{text Var} is now +called @{term Atom}. + 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) +consts valid :: "state \ formula \ bool" ("(_ \ _)" [80,80] 80) + +text{*\noindent +The concrete syntax annotation allows us to write @{term"s \ f"} instead of +@{text"valid s f"}. + +The definition of @{text"\"} is by recursion over the syntax: +*} primrec "s \ Atom a = (a \ L s)" -"s \ NOT f = (\(s \ f))" +"s \ Neg 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)"; -text{* +text{*\noindent +The first three equations should be self-explanatory. The temporal formula +@{term"AX f"} means that @{term f} is true in all next states whereas +@{term"EF f"} means that there exists some future state in which @{term f} is +true. The future is expressed via @{text"^*"}, the transitive reflexive +closure. Because of reflexivity, the future includes the present. + 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: +states where the formula is true and is defined by recursion over the syntax, +too: *} -consts mc :: "pdl_form \ state set"; +consts mc :: "formula \ state set"; primrec "mc(Atom a) = {s. a \ L s}" -"mc(NOT f) = -mc f" +"mc(Neg f) = -mc f" "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 \ M^-1 ^^ 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}. +text{*\noindent +Only the equation for @{term EF} deserves some comments. Remember that 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 the least +fixpoint (@{term lfp}) of @{term"\T. mc f \ M^-1 ^^ T"} is the least set +@{term T} containing @{term"mc f"} and all predecessors of @{term T}. If you +find it hard to see that @{term"mc(EF f)"} contains exactly those states from +which there is a path to a state where @{term f} is true, do not worry---that +will be proved in a moment. + +First we prove monotonicity of the function inside @{term lfp} *} -lemma mono_lemma: "mono(\T. A \ M^-1 ^^ T)" -apply(rule monoI); -by(blast); +lemma mono_ef: "mono(\T. A \ M^-1 ^^ T)" +apply(rule monoI) +by(blast) + +text{*\noindent +in order to make sure it really has a least fixpoint. -lemma lfp_conv_EF: -"lfp(\T. A \ M^-1 ^^ T) = {s. \t. (s,t) \ M^* \ t \ A}"; +Now we can relate model checking and semantics. For the @{text EF} case we need +a separate lemma: +*} + +lemma EF_lemma: + "lfp(\T. A \ M^-1 ^^ T) = {s. \t. (s,t) \ M^* \ t \ A}" + +txt{*\noindent +The equality is proved in the canonical fashion by proving that each set +contains the other; the containment is shown pointwise: +*} + apply(rule equalityI); apply(rule subsetI); - apply(simp); - apply(erule Lfp.induct); - apply(rule mono_lemma); - apply(simp); + apply(simp) +(*pr(latex xsymbols symbols);*) +txt{*\noindent +Simplification leaves us with the following first subgoal +\begin{isabelle} +\ \isadigit{1}{\isachardot}\ {\isasymAnd}s{\isachardot}\ s\ {\isasymin}\ lfp\ {\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ T{\isacharparenright}\ {\isasymLongrightarrow}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A +\end{isabelle} +which is proved by @{term lfp}-induction: +*} + + apply(erule Lfp.induct) + apply(rule mono_ef) + apply(simp) +(*pr(latex xsymbols symbols);*) +txt{*\noindent +Having disposed of the monotonicity subgoal, +simplification leaves us with the following main goal +\begin{isabelle} +\ \isadigit{1}{\isachardot}\ {\isasymAnd}s{\isachardot}\ s\ {\isasymin}\ A\ {\isasymor}\isanewline +\ \ \ \ \ \ \ \ \ s\ {\isasymin}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ {\isacharparenleft}lfp\ {\isacharparenleft}{\dots}{\isacharparenright}\ {\isasyminter}\ {\isacharbraceleft}x{\isachardot}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}x{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A{\isacharbraceright}{\isacharparenright}\isanewline +\ \ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A +\end{isabelle} +which is proved by @{text blast} with the help of a few lemmas about +@{text"^*"}: +*} + apply(blast intro: r_into_rtrancl rtrancl_trans); -apply(rule subsetI); -apply(simp); -apply(erule exE); -apply(erule conjE); -apply(erule_tac P = "t\A" in rev_mp); -apply(erule converse_rtrancl_induct); - apply(rule ssubst[OF lfp_Tarski[OF mono_lemma]]); - apply(blast); -apply(rule ssubst[OF lfp_Tarski[OF mono_lemma]]); -by(blast); + +txt{* +We now return to the second set containment subgoal, which is again proved +pointwise: +*} + +apply(rule subsetI) +apply(simp, clarify) +(*pr(latex xsymbols symbols);*) +txt{*\noindent +After simplification and clarification we are left with +\begin{isabelle} +\ \isadigit{1}{\isachardot}\ {\isasymAnd}s\ t{\isachardot}\ {\isasymlbrakk}{\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}{\isacharsemicolon}\ t\ {\isasymin}\ A{\isasymrbrakk}\ {\isasymLongrightarrow}\ s\ {\isasymin}\ lfp\ {\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ T{\isacharparenright} +\end{isabelle} +This goal is proved by induction on @{term"(s,t)\M^*"}. But since the model +checker works backwards (from @{term t} to @{term s}), we cannot use the +induction theorem @{thm[source]rtrancl_induct} because it works in the +forward direction. Fortunately the converse induction theorem +@{thm[source]converse_rtrancl_induct} already exists: +@{thm[display,margin=60]converse_rtrancl_induct[no_vars]} +It says that if @{prop"(a,b):r^*"} and we know @{prop"P b"} then we can infer +@{prop"P a"} provided each step backwards from a predecessor @{term z} of +@{term b} preserves @{term P}. +*} + +apply(erule converse_rtrancl_induct) +(*pr(latex xsymbols symbols);*) +txt{*\noindent +The base case +\begin{isabelle} +\ \isadigit{1}{\isachardot}\ {\isasymAnd}t{\isachardot}\ t\ {\isasymin}\ A\ {\isasymLongrightarrow}\ t\ {\isasymin}\ lfp\ {\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ T{\isacharparenright} +\end{isabelle} +is solved by unrolling @{term lfp} once +*} + + apply(rule ssubst[OF lfp_Tarski[OF mono_ef]]) +(*pr(latex xsymbols symbols);*) +txt{* +\begin{isabelle} +\ \isadigit{1}{\isachardot}\ {\isasymAnd}t{\isachardot}\ t\ {\isasymin}\ A\ {\isasymLongrightarrow}\ t\ {\isasymin}\ A\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ lfp\ {\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ T{\isacharparenright} +\end{isabelle} +and disposing of the resulting trivial subgoal automatically: +*} + + apply(blast) + +txt{*\noindent +The proof of the induction step is identical to the one for the base case: +*} + +apply(rule ssubst[OF lfp_Tarski[OF mono_ef]]) +by(blast) + +text{* +The main theorem is proved in the familiar manner: induction followed by +@{text auto} augmented with the lemma as a simplification rule. +*} theorem "mc f = {s. s \ f}"; apply(induct_tac f); -by(auto simp add:lfp_conv_EF); +by(auto simp add:EF_lemma); (*<*)end(*>*) diff -r 739327964a5c -r 7cfdf6a330a0 doc-src/TutorialI/CTL/document/CTL.tex --- a/doc-src/TutorialI/CTL/document/CTL.tex Tue Oct 03 22:39:49 2000 +0200 +++ b/doc-src/TutorialI/CTL/document/CTL.tex Wed Oct 04 17:35:45 2000 +0200 @@ -3,65 +3,51 @@ \def\isabellecontext{CTL}% % \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 -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ AX\ ctl{\isacharunderscore}form\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ EF\ ctl{\isacharunderscore}form\isanewline -\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\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 -\isanewline +% +\begin{isamarkuptext}% +The semantics of PDL only needs transitive reflexive closure. +Let us now be a bit more adventurous and introduce a new temporal operator +that goes beyond transitive reflexive closure. We extend the datatype +\isa{formula} by a new constructor% +\end{isamarkuptext}% +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ AF\ formula% +\begin{isamarkuptext}% +\noindent +which stands for "always in the future": +on all paths, at some point the formula holds. +Introducing the notion of paths (in \isa{M})% +\end{isamarkuptext}% \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}\ 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 -{\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 -{\isachardoublequote}s\ {\isasymTurnstile}\ AF\ f\ \ \ \ {\isacharequal}\ {\isacharparenleft}{\isasymforall}p\ {\isasymin}\ Paths\ s{\isachardot}\ {\isasymexists}i{\isachardot}\ p\ i\ {\isasymTurnstile}\ f{\isacharparenright}{\isachardoublequote}\isanewline -\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}% +\begin{isamarkuptext}% +\noindent +allows a very succinct definition of the semantics of \isa{AF}: +\footnote{Do not be mislead: neither datatypes nor recursive functions can be +extended by new constructors or equations. This is just a trick of the +presentation. In reality one has to define a new datatype and a new function.}% +\end{isamarkuptext}% +{\isachardoublequote}s\ {\isasymTurnstile}\ AF\ f\ \ \ \ {\isacharequal}\ {\isacharparenleft}{\isasymforall}p\ {\isasymin}\ Paths\ s{\isachardot}\ {\isasymexists}i{\isachardot}\ p\ i\ {\isasymTurnstile}\ f{\isacharparenright}{\isachardoublequote}% +\begin{isamarkuptext}% +\noindent +Model checking \isa{AF} involves a function which +is just large enough to warrant a separate definition:% +\end{isamarkuptext}% \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}\isanewline -\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}% +\begin{isamarkuptext}% +\noindent +This function is monotone in its second argument (and also its first, but +that is irrelevant), and hence \isa{af\ A} has a least fixpoint.% +\end{isamarkuptext}% \isacommand{lemma}\ mono{\isacharunderscore}af{\isacharcolon}\ {\isachardoublequote}mono{\isacharparenleft}af\ A{\isacharparenright}{\isachardoublequote}\isanewline -\isacommand{by}{\isacharparenleft}force\ simp\ add{\isacharcolon}\ af{\isacharunderscore}def\ intro{\isacharcolon}monoI{\isacharparenright}\isanewline -\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}\ 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 -{\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 +\isacommand{apply}{\isacharparenleft}simp\ add{\isacharcolon}\ mono{\isacharunderscore}def\ af{\isacharunderscore}def{\isacharparenright}\isanewline +\isacommand{by}{\isacharparenleft}blast{\isacharparenright}% +\begin{isamarkuptext}% +\noindent +Now we can define \isa{mc\ {\isacharparenleft}AF\ f{\isacharparenright}} as the least set \isa{T} that contains +\isa{mc\ f} and all states all of whose direct successors are in \isa{T}:% +\end{isamarkuptext}% {\isachardoublequote}mc{\isacharparenleft}AF\ f{\isacharparenright}\ \ \ \ {\isacharequal}\ lfp{\isacharparenleft}af{\isacharparenleft}mc\ f{\isacharparenright}{\isacharparenright}{\isachardoublequote}\isanewline -\isanewline -\isacommand{lemma}\ mono{\isacharunderscore}ef{\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 -\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 -\isacommand{apply}{\isacharparenleft}rule\ equalityI{\isacharparenright}\isanewline -\ \isacommand{apply}{\isacharparenleft}rule\ subsetI{\isacharparenright}\isanewline -\ \isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isanewline -\ \isacommand{apply}{\isacharparenleft}erule\ Lfp{\isachardot}induct{\isacharparenright}\isanewline -\ \ \isacommand{apply}{\isacharparenleft}rule\ mono{\isacharunderscore}ef{\isacharparenright}\isanewline -\ \isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isanewline -\ \isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ r{\isacharunderscore}into{\isacharunderscore}rtrancl\ rtrancl{\isacharunderscore}trans{\isacharparenright}\isanewline -\isacommand{apply}{\isacharparenleft}rule\ subsetI{\isacharparenright}\isanewline -\isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isanewline -\isacommand{apply}{\isacharparenleft}erule\ exE{\isacharparenright}\isanewline -\isacommand{apply}{\isacharparenleft}erule\ conjE{\isacharparenright}\isanewline -\isacommand{apply}{\isacharparenleft}erule{\isacharunderscore}tac\ P\ {\isacharequal}\ {\isachardoublequote}t{\isasymin}A{\isachardoublequote}\ \isakeyword{in}\ rev{\isacharunderscore}mp{\isacharparenright}\isanewline -\isacommand{apply}{\isacharparenleft}erule\ converse{\isacharunderscore}rtrancl{\isacharunderscore}induct{\isacharparenright}\isanewline -\ \isacommand{apply}{\isacharparenleft}rule\ ssubst\ {\isacharbrackleft}OF\ lfp{\isacharunderscore}Tarski{\isacharbrackleft}OF\ mono{\isacharunderscore}ef{\isacharbrackright}{\isacharbrackright}{\isacharparenright}\isanewline -\ \isacommand{apply}{\isacharparenleft}blast{\isacharparenright}\isanewline -\isacommand{apply}{\isacharparenleft}rule\ ssubst\ {\isacharbrackleft}OF\ lfp{\isacharunderscore}Tarski{\isacharbrackleft}OF\ mono{\isacharunderscore}ef{\isacharbrackright}{\isacharbrackright}{\isacharparenright}\isanewline -\isacommand{by}{\isacharparenleft}blast{\isacharparenright}\isanewline -\isanewline \isacommand{theorem}\ lfp{\isacharunderscore}subset{\isacharunderscore}AF{\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}\isanewline \isacommand{apply}{\isacharparenleft}rule\ subsetI{\isacharparenright}\isanewline @@ -214,7 +200,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\ equalityI{\isacharbrackleft}OF\ lfp{\isacharunderscore}subset{\isacharunderscore}AF\ AF{\isacharunderscore}subset{\isacharunderscore}lfp{\isacharbrackright}{\isacharparenright}\isanewline +\isacommand{by}{\isacharparenleft}auto\ simp\ add{\isacharcolon}\ EF{\isacharunderscore}lemma\ equalityI{\isacharbrackleft}OF\ lfp{\isacharunderscore}subset{\isacharunderscore}AF\ AF{\isacharunderscore}subset{\isacharunderscore}lfp{\isacharbrackright}{\isacharparenright}\isanewline \end{isabellebody}% %%% Local Variables: %%% mode: latex diff -r 739327964a5c -r 7cfdf6a330a0 doc-src/TutorialI/CTL/document/PDL.tex --- a/doc-src/TutorialI/CTL/document/PDL.tex Tue Oct 03 22:39:49 2000 +0200 +++ b/doc-src/TutorialI/CTL/document/PDL.tex Wed Oct 04 17:35:45 2000 +0200 @@ -10,69 +10,167 @@ 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% +\isacommand{datatype}\ formula\ {\isacharequal}\ Atom\ atom\isanewline +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ Neg\ formula\isanewline +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ And\ formula\ formula\isanewline +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ AX\ formula\isanewline +\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ EF\ formula% \begin{isamarkuptext}% \noindent +This is almost the same as in the boolean expression case study in +\S\ref{sec:boolex}, except that what used to be called \isa{Var} is now +called \isa{formula{\isachardot}Atom}. + 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{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}% +\begin{isamarkuptext}% +\noindent +The concrete syntax annotation allows us to write \isa{s\ {\isasymTurnstile}\ f} instead of +\isa{valid\ s\ f}. + +The definition of \isa{{\isasymTurnstile}} is by recursion over the syntax:% +\end{isamarkuptext}% \isacommand{primrec}\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}\ 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{\isacharcircum}{\isacharasterisk}\ {\isasymand}\ t\ {\isasymTurnstile}\ f{\isacharparenright}{\isachardoublequote}% \begin{isamarkuptext}% +\noindent +The first three equations should be self-explanatory. The temporal formula +\isa{AX\ f} means that \isa{f} is true in all next states whereas +\isa{EF\ f} means that there exists some future state in which \isa{f} is +true. The future is expressed via \isa{{\isacharcircum}{\isacharasterisk}}, the transitive reflexive +closure. Because of reflexivity, the future includes the present. + 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:% +states where the formula is true and is defined by recursion over the syntax, +too:% \end{isamarkuptext}% -\isacommand{consts}\ mc\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}pdl{\isacharunderscore}form\ {\isasymRightarrow}\ state\ set{\isachardoublequote}\isanewline +\isacommand{consts}\ mc\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}formula\ {\isasymRightarrow}\ state\ set{\isachardoublequote}\isanewline \isacommand{primrec}\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}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}\ 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}.% +\noindent +Only the equation for \isa{EF} deserves some comments. Remember that 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 the least +fixpoint (\isa{lfp}) of \isa{{\isasymlambda}T{\isachardot}\ mc\ f\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ T} is the least set +\isa{T} containing \isa{mc\ f} and all predecessors of \isa{T}. If you +find it hard to see that \isa{mc\ {\isacharparenleft}EF\ f{\isacharparenright}} contains exactly those states from +which there is a path to a state where \isa{f} is true, do not worry---that +will be proved in a moment. + +First we prove monotonicity of the function inside \isa{lfp}% \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{lemma}\ mono{\isacharunderscore}ef{\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}\ 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{by}{\isacharparenleft}blast{\isacharparenright}% +\begin{isamarkuptext}% +\noindent +in order to make sure it really has a least fixpoint. + +Now we can relate model checking and semantics. For the \isa{EF} case we need +a separate lemma:% +\end{isamarkuptext}% +\isacommand{lemma}\ EF{\isacharunderscore}lemma{\isacharcolon}\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}% +\begin{isamarkuptxt}% +\noindent +The equality is proved in the canonical fashion by proving that each set +contains the other; the containment is shown pointwise:% +\end{isamarkuptxt}% \isacommand{apply}{\isacharparenleft}rule\ equalityI{\isacharparenright}\isanewline \ \isacommand{apply}{\isacharparenleft}rule\ subsetI{\isacharparenright}\isanewline -\ \isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isanewline +\ \isacommand{apply}{\isacharparenleft}simp{\isacharparenright}% +\begin{isamarkuptxt}% +\noindent +Simplification leaves us with the following first subgoal +\begin{isabelle} +\ \isadigit{1}{\isachardot}\ {\isasymAnd}s{\isachardot}\ s\ {\isasymin}\ lfp\ {\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ T{\isacharparenright}\ {\isasymLongrightarrow}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A +\end{isabelle} +which is proved by \isa{lfp}-induction:% +\end{isamarkuptxt}% \ \isacommand{apply}{\isacharparenleft}erule\ Lfp{\isachardot}induct{\isacharparenright}\isanewline -\ \ \isacommand{apply}{\isacharparenleft}rule\ mono{\isacharunderscore}lemma{\isacharparenright}\isanewline -\ \isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isanewline -\ \isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ r{\isacharunderscore}into{\isacharunderscore}rtrancl\ rtrancl{\isacharunderscore}trans{\isacharparenright}\isanewline +\ \ \isacommand{apply}{\isacharparenleft}rule\ mono{\isacharunderscore}ef{\isacharparenright}\isanewline +\ \isacommand{apply}{\isacharparenleft}simp{\isacharparenright}% +\begin{isamarkuptxt}% +\noindent +Having disposed of the monotonicity subgoal, +simplification leaves us with the following main goal +\begin{isabelle} +\ \isadigit{1}{\isachardot}\ {\isasymAnd}s{\isachardot}\ s\ {\isasymin}\ A\ {\isasymor}\isanewline +\ \ \ \ \ \ \ \ \ s\ {\isasymin}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ {\isacharparenleft}lfp\ {\isacharparenleft}{\dots}{\isacharparenright}\ {\isasyminter}\ {\isacharbraceleft}x{\isachardot}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}x{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A{\isacharbraceright}{\isacharparenright}\isanewline +\ \ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A +\end{isabelle} +which is proved by \isa{blast} with the help of a few lemmas about +\isa{{\isacharcircum}{\isacharasterisk}}:% +\end{isamarkuptxt}% +\ \isacommand{apply}{\isacharparenleft}blast\ intro{\isacharcolon}\ r{\isacharunderscore}into{\isacharunderscore}rtrancl\ rtrancl{\isacharunderscore}trans{\isacharparenright}% +\begin{isamarkuptxt}% +We now return to the second set containment subgoal, which is again proved +pointwise:% +\end{isamarkuptxt}% \isacommand{apply}{\isacharparenleft}rule\ subsetI{\isacharparenright}\isanewline -\isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isanewline -\isacommand{apply}{\isacharparenleft}erule\ exE{\isacharparenright}\isanewline -\isacommand{apply}{\isacharparenleft}erule\ conjE{\isacharparenright}\isanewline -\isacommand{apply}{\isacharparenleft}erule{\isacharunderscore}tac\ P\ {\isacharequal}\ {\isachardoublequote}t{\isasymin}A{\isachardoublequote}\ \isakeyword{in}\ rev{\isacharunderscore}mp{\isacharparenright}\isanewline -\isacommand{apply}{\isacharparenleft}erule\ converse{\isacharunderscore}rtrancl{\isacharunderscore}induct{\isacharparenright}\isanewline -\ \isacommand{apply}{\isacharparenleft}rule\ ssubst{\isacharbrackleft}OF\ lfp{\isacharunderscore}Tarski{\isacharbrackleft}OF\ mono{\isacharunderscore}lemma{\isacharbrackright}{\isacharbrackright}{\isacharparenright}\isanewline -\ \isacommand{apply}{\isacharparenleft}blast{\isacharparenright}\isanewline -\isacommand{apply}{\isacharparenleft}rule\ ssubst{\isacharbrackleft}OF\ lfp{\isacharunderscore}Tarski{\isacharbrackleft}OF\ mono{\isacharunderscore}lemma{\isacharbrackright}{\isacharbrackright}{\isacharparenright}\isanewline -\isacommand{by}{\isacharparenleft}blast{\isacharparenright}\isanewline -\isanewline +\isacommand{apply}{\isacharparenleft}simp{\isacharcomma}\ clarify{\isacharparenright}% +\begin{isamarkuptxt}% +\noindent +After simplification and clarification we are left with +\begin{isabelle} +\ \isadigit{1}{\isachardot}\ {\isasymAnd}s\ t{\isachardot}\ {\isasymlbrakk}{\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}{\isacharsemicolon}\ t\ {\isasymin}\ A{\isasymrbrakk}\ {\isasymLongrightarrow}\ s\ {\isasymin}\ lfp\ {\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ T{\isacharparenright} +\end{isabelle} +This goal is proved by induction on \isa{{\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}}. But since the model +checker works backwards (from \isa{t} to \isa{s}), we cannot use the +induction theorem \isa{rtrancl{\isacharunderscore}induct} because it works in the +forward direction. Fortunately the converse induction theorem +\isa{converse{\isacharunderscore}rtrancl{\isacharunderscore}induct} already exists: +\begin{isabelle}% +\ \ \ \ \ {\isasymlbrakk}{\isacharparenleft}a{\isacharcomma}\ b{\isacharparenright}\ {\isasymin}\ r{\isacharcircum}{\isacharasterisk}{\isacharsemicolon}\ P\ b{\isacharsemicolon}\isanewline +\ \ \ \ \ \ \ \ {\isasymAnd}y\ z{\isachardot}\ {\isasymlbrakk}{\isacharparenleft}y{\isacharcomma}\ z{\isacharparenright}\ {\isasymin}\ r{\isacharsemicolon}\ {\isacharparenleft}z{\isacharcomma}\ b{\isacharparenright}\ {\isasymin}\ r{\isacharcircum}{\isacharasterisk}{\isacharsemicolon}\ P\ z{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ y{\isasymrbrakk}\isanewline +\ \ \ \ \ {\isasymLongrightarrow}\ P\ a% +\end{isabelle} +It says that if \isa{{\isacharparenleft}a{\isacharcomma}\ b{\isacharparenright}\ {\isasymin}\ r{\isacharcircum}{\isacharasterisk}} and we know \isa{P\ b} then we can infer +\isa{P\ a} provided each step backwards from a predecessor \isa{z} of +\isa{b} preserves \isa{P}.% +\end{isamarkuptxt}% +\isacommand{apply}{\isacharparenleft}erule\ converse{\isacharunderscore}rtrancl{\isacharunderscore}induct{\isacharparenright}% +\begin{isamarkuptxt}% +\noindent +The base case +\begin{isabelle} +\ \isadigit{1}{\isachardot}\ {\isasymAnd}t{\isachardot}\ t\ {\isasymin}\ A\ {\isasymLongrightarrow}\ t\ {\isasymin}\ lfp\ {\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ T{\isacharparenright} +\end{isabelle} +is solved by unrolling \isa{lfp} once% +\end{isamarkuptxt}% +\ \isacommand{apply}{\isacharparenleft}rule\ ssubst{\isacharbrackleft}OF\ lfp{\isacharunderscore}Tarski{\isacharbrackleft}OF\ mono{\isacharunderscore}ef{\isacharbrackright}{\isacharbrackright}{\isacharparenright}% +\begin{isamarkuptxt}% +\begin{isabelle} +\ \isadigit{1}{\isachardot}\ {\isasymAnd}t{\isachardot}\ t\ {\isasymin}\ A\ {\isasymLongrightarrow}\ t\ {\isasymin}\ A\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ lfp\ {\isacharparenleft}{\isasymlambda}T{\isachardot}\ A\ {\isasymunion}\ M{\isacharcircum}{\isacharminus}\isadigit{1}\ {\isacharcircum}{\isacharcircum}\ T{\isacharparenright} +\end{isabelle} +and disposing of the resulting trivial subgoal automatically:% +\end{isamarkuptxt}% +\ \isacommand{apply}{\isacharparenleft}blast{\isacharparenright}% +\begin{isamarkuptxt}% +\noindent +The proof of the induction step is identical to the one for the base case:% +\end{isamarkuptxt}% +\isacommand{apply}{\isacharparenleft}rule\ ssubst{\isacharbrackleft}OF\ lfp{\isacharunderscore}Tarski{\isacharbrackleft}OF\ mono{\isacharunderscore}ef{\isacharbrackright}{\isacharbrackright}{\isacharparenright}\isanewline +\isacommand{by}{\isacharparenleft}blast{\isacharparenright}% +\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}% \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}\end{isabellebody}% +\isacommand{by}{\isacharparenleft}auto\ simp\ add{\isacharcolon}EF{\isacharunderscore}lemma{\isacharparenright}\end{isabellebody}% %%% Local Variables: %%% mode: latex %%% TeX-master: "root"