doc-src/TutorialI/CTL/PDL.thy
changeset 10149 7cfdf6a330a0
parent 10133 e187dacd248f
child 10159 a72ddfdbfca0
     1.1 --- a/doc-src/TutorialI/CTL/PDL.thy	Tue Oct 03 22:39:49 2000 +0200
     1.2 +++ b/doc-src/TutorialI/CTL/PDL.thy	Wed Oct 04 17:35:45 2000 +0200
     1.3 @@ -9,73 +9,179 @@
     1.4  (syntax) trees, they are naturally modelled as a datatype:
     1.5  *}
     1.6  
     1.7 -datatype pdl_form = Atom atom
     1.8 -                  | NOT pdl_form
     1.9 -                  | And pdl_form pdl_form
    1.10 -                  | AX pdl_form
    1.11 -                  | EF pdl_form;
    1.12 +datatype formula = Atom atom
    1.13 +                  | Neg formula
    1.14 +                  | And formula formula
    1.15 +                  | AX formula
    1.16 +                  | EF formula
    1.17  
    1.18  text{*\noindent
    1.19 +This is almost the same as in the boolean expression case study in
    1.20 +\S\ref{sec:boolex}, except that what used to be called @{text Var} is now
    1.21 +called @{term Atom}.
    1.22 +
    1.23  The meaning of these formulae is given by saying which formula is true in
    1.24  which state:
    1.25  *}
    1.26  
    1.27 -consts valid :: "state \<Rightarrow> pdl_form \<Rightarrow> bool" ("(_ \<Turnstile> _)" [80,80] 80)
    1.28 +consts valid :: "state \<Rightarrow> formula \<Rightarrow> bool"   ("(_ \<Turnstile> _)" [80,80] 80)
    1.29 +
    1.30 +text{*\noindent
    1.31 +The concrete syntax annotation allows us to write @{term"s \<Turnstile> f"} instead of
    1.32 +@{text"valid s f"}.
    1.33 +
    1.34 +The definition of @{text"\<Turnstile>"} is by recursion over the syntax:
    1.35 +*}
    1.36  
    1.37  primrec
    1.38  "s \<Turnstile> Atom a  = (a \<in> L s)"
    1.39 -"s \<Turnstile> NOT f   = (\<not>(s \<Turnstile> f))"
    1.40 +"s \<Turnstile> Neg f   = (\<not>(s \<Turnstile> f))"
    1.41  "s \<Turnstile> And f g = (s \<Turnstile> f \<and> s \<Turnstile> g)"
    1.42  "s \<Turnstile> AX f    = (\<forall>t. (s,t) \<in> M \<longrightarrow> t \<Turnstile> f)"
    1.43  "s \<Turnstile> EF f    = (\<exists>t. (s,t) \<in> M^* \<and> t \<Turnstile> f)";
    1.44  
    1.45 -text{*
    1.46 +text{*\noindent
    1.47 +The first three equations should be self-explanatory. The temporal formula
    1.48 +@{term"AX f"} means that @{term f} is true in all next states whereas
    1.49 +@{term"EF f"} means that there exists some future state in which @{term f} is
    1.50 +true. The future is expressed via @{text"^*"}, the transitive reflexive
    1.51 +closure. Because of reflexivity, the future includes the present.
    1.52 +
    1.53  Now we come to the model checker itself. It maps a formula into the set of
    1.54 -states where the formula is true and is defined by recursion over the syntax:
    1.55 +states where the formula is true and is defined by recursion over the syntax,
    1.56 +too:
    1.57  *}
    1.58  
    1.59 -consts mc :: "pdl_form \<Rightarrow> state set";
    1.60 +consts mc :: "formula \<Rightarrow> state set";
    1.61  primrec
    1.62  "mc(Atom a)  = {s. a \<in> L s}"
    1.63 -"mc(NOT f)   = -mc f"
    1.64 +"mc(Neg f)   = -mc f"
    1.65  "mc(And f g) = mc f \<inter> mc g"
    1.66  "mc(AX f)    = {s. \<forall>t. (s,t) \<in> M  \<longrightarrow> t \<in> mc f}"
    1.67  "mc(EF f)    = lfp(\<lambda>T. mc f \<union> M^-1 ^^ T)"
    1.68  
    1.69  
    1.70 -text{*
    1.71 -Only the equation for @{term EF} deserves a comment: the postfix @{text"^-1"}
    1.72 -and the infix @{text"^^"} are predefined and denote the converse of a
    1.73 -relation and the application of a relation to a set. Thus @{term "M^-1 ^^ T"}
    1.74 -is the set of all predecessors of @{term T} and @{term"mc(EF f)"} is the least
    1.75 -set @{term T} containing @{term"mc f"} and all predecessors of @{term T}.
    1.76 +text{*\noindent
    1.77 +Only the equation for @{term EF} deserves some comments. Remember that the
    1.78 +postfix @{text"^-1"} and the infix @{text"^^"} are predefined and denote the
    1.79 +converse of a relation and the application of a relation to a set. Thus
    1.80 +@{term "M^-1 ^^ T"} is the set of all predecessors of @{term T} and the least
    1.81 +fixpoint (@{term lfp}) of @{term"\<lambda>T. mc f \<union> M^-1 ^^ T"} is the least set
    1.82 +@{term T} containing @{term"mc f"} and all predecessors of @{term T}. If you
    1.83 +find it hard to see that @{term"mc(EF f)"} contains exactly those states from
    1.84 +which there is a path to a state where @{term f} is true, do not worry---that
    1.85 +will be proved in a moment.
    1.86 +
    1.87 +First we prove monotonicity of the function inside @{term lfp}
    1.88  *}
    1.89  
    1.90 -lemma mono_lemma: "mono(\<lambda>T. A \<union> M^-1 ^^ T)"
    1.91 -apply(rule monoI);
    1.92 -by(blast);
    1.93 +lemma mono_ef: "mono(\<lambda>T. A \<union> M^-1 ^^ T)"
    1.94 +apply(rule monoI)
    1.95 +by(blast)
    1.96  
    1.97 -lemma lfp_conv_EF:
    1.98 -"lfp(\<lambda>T. A \<union> M^-1 ^^ T) = {s. \<exists>t. (s,t) \<in> M^* \<and> t \<in> A}";
    1.99 +text{*\noindent
   1.100 +in order to make sure it really has a least fixpoint.
   1.101 +
   1.102 +Now we can relate model checking and semantics. For the @{text EF} case we need
   1.103 +a separate lemma:
   1.104 +*}
   1.105 +
   1.106 +lemma EF_lemma:
   1.107 +  "lfp(\<lambda>T. A \<union> M^-1 ^^ T) = {s. \<exists>t. (s,t) \<in> M^* \<and> t \<in> A}"
   1.108 +
   1.109 +txt{*\noindent
   1.110 +The equality is proved in the canonical fashion by proving that each set
   1.111 +contains the other; the containment is shown pointwise:
   1.112 +*}
   1.113 +
   1.114  apply(rule equalityI);
   1.115   apply(rule subsetI);
   1.116 - apply(simp);
   1.117 - apply(erule Lfp.induct);
   1.118 -  apply(rule mono_lemma);
   1.119 - apply(simp);
   1.120 + apply(simp)
   1.121 +(*pr(latex xsymbols symbols);*)
   1.122 +txt{*\noindent
   1.123 +Simplification leaves us with the following first subgoal
   1.124 +\begin{isabelle}
   1.125 +\ \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
   1.126 +\end{isabelle}
   1.127 +which is proved by @{term lfp}-induction:
   1.128 +*}
   1.129 +
   1.130 + apply(erule Lfp.induct)
   1.131 +  apply(rule mono_ef)
   1.132 + apply(simp)
   1.133 +(*pr(latex xsymbols symbols);*)
   1.134 +txt{*\noindent
   1.135 +Having disposed of the monotonicity subgoal,
   1.136 +simplification leaves us with the following main goal
   1.137 +\begin{isabelle}
   1.138 +\ \isadigit{1}{\isachardot}\ {\isasymAnd}s{\isachardot}\ s\ {\isasymin}\ A\ {\isasymor}\isanewline
   1.139 +\ \ \ \ \ \ \ \ \ 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
   1.140 +\ \ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}s{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M{\isacharcircum}{\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A
   1.141 +\end{isabelle}
   1.142 +which is proved by @{text blast} with the help of a few lemmas about
   1.143 +@{text"^*"}:
   1.144 +*}
   1.145 +
   1.146   apply(blast intro: r_into_rtrancl rtrancl_trans);
   1.147 -apply(rule subsetI);
   1.148 -apply(simp);
   1.149 -apply(erule exE);
   1.150 -apply(erule conjE);
   1.151 -apply(erule_tac P = "t\<in>A" in rev_mp);
   1.152 -apply(erule converse_rtrancl_induct);
   1.153 - apply(rule ssubst[OF lfp_Tarski[OF mono_lemma]]);
   1.154 - apply(blast);
   1.155 -apply(rule ssubst[OF lfp_Tarski[OF mono_lemma]]);
   1.156 -by(blast);
   1.157 +
   1.158 +txt{*
   1.159 +We now return to the second set containment subgoal, which is again proved
   1.160 +pointwise:
   1.161 +*}
   1.162 +
   1.163 +apply(rule subsetI)
   1.164 +apply(simp, clarify)
   1.165 +(*pr(latex xsymbols symbols);*)
   1.166 +txt{*\noindent
   1.167 +After simplification and clarification we are left with
   1.168 +\begin{isabelle}
   1.169 +\ \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}
   1.170 +\end{isabelle}
   1.171 +This goal is proved by induction on @{term"(s,t)\<in>M^*"}. But since the model
   1.172 +checker works backwards (from @{term t} to @{term s}), we cannot use the
   1.173 +induction theorem @{thm[source]rtrancl_induct} because it works in the
   1.174 +forward direction. Fortunately the converse induction theorem
   1.175 +@{thm[source]converse_rtrancl_induct} already exists:
   1.176 +@{thm[display,margin=60]converse_rtrancl_induct[no_vars]}
   1.177 +It says that if @{prop"(a,b):r^*"} and we know @{prop"P b"} then we can infer
   1.178 +@{prop"P a"} provided each step backwards from a predecessor @{term z} of
   1.179 +@{term b} preserves @{term P}.
   1.180 +*}
   1.181 +
   1.182 +apply(erule converse_rtrancl_induct)
   1.183 +(*pr(latex xsymbols symbols);*)
   1.184 +txt{*\noindent
   1.185 +The base case
   1.186 +\begin{isabelle}
   1.187 +\ \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}
   1.188 +\end{isabelle}
   1.189 +is solved by unrolling @{term lfp} once
   1.190 +*}
   1.191 +
   1.192 + apply(rule ssubst[OF lfp_Tarski[OF mono_ef]])
   1.193 +(*pr(latex xsymbols symbols);*)
   1.194 +txt{*
   1.195 +\begin{isabelle}
   1.196 +\ \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}
   1.197 +\end{isabelle}
   1.198 +and disposing of the resulting trivial subgoal automatically:
   1.199 +*}
   1.200 +
   1.201 + apply(blast)
   1.202 +
   1.203 +txt{*\noindent
   1.204 +The proof of the induction step is identical to the one for the base case:
   1.205 +*}
   1.206 +
   1.207 +apply(rule ssubst[OF lfp_Tarski[OF mono_ef]])
   1.208 +by(blast)
   1.209 +
   1.210 +text{*
   1.211 +The main theorem is proved in the familiar manner: induction followed by
   1.212 +@{text auto} augmented with the lemma as a simplification rule.
   1.213 +*}
   1.214  
   1.215  theorem "mc f = {s. s \<Turnstile> f}";
   1.216  apply(induct_tac f);
   1.217 -by(auto simp add:lfp_conv_EF);
   1.218 +by(auto simp add:EF_lemma);
   1.219  (*<*)end(*>*)