doc-src/TutorialI/CTL/PDL.thy
author nipkow
Thu, 29 May 2008 22:45:33 +0200
changeset 27015 f8537d69f514
parent 21202 6649bf75b9dc
permissions -rw-r--r--
*** empty log message ***
     1 (*<*)theory PDL imports Base begin(*>*)
     2 
     3 subsection{*Propositional Dynamic Logic --- PDL*}
     4 
     5 text{*\index{PDL|(}
     6 The formulae of PDL are built up from atomic propositions via
     7 negation and conjunction and the two temporal
     8 connectives @{text AX} and @{text EF}\@. Since formulae are essentially
     9 syntax trees, they are naturally modelled as a datatype:%
    10 \footnote{The customary definition of PDL
    11 \cite{HarelKT-DL} looks quite different from ours, but the two are easily
    12 shown to be equivalent.}
    13 *}
    14 
    15 datatype formula = Atom "atom"
    16                   | Neg formula
    17                   | And formula formula
    18                   | AX formula
    19                   | EF formula
    20 
    21 text{*\noindent
    22 This resembles the boolean expression case study in
    23 \S\ref{sec:boolex}.
    24 A validity relation between states and formulae specifies the semantics.
    25 The syntax annotation allows us to write @{text"s \<Turnstile> f"} instead of
    26 \hbox{@{text"valid s f"}}. The definition is by recursion over the syntax:
    27 *}
    28 
    29 primrec valid :: "state \<Rightarrow> formula \<Rightarrow> bool"   ("(_ \<Turnstile> _)" [80,80] 80)
    30 where
    31 "s \<Turnstile> Atom a  = (a \<in> L s)" |
    32 "s \<Turnstile> Neg f   = (\<not>(s \<Turnstile> f))" |
    33 "s \<Turnstile> And f g = (s \<Turnstile> f \<and> s \<Turnstile> g)" |
    34 "s \<Turnstile> AX f    = (\<forall>t. (s,t) \<in> M \<longrightarrow> t \<Turnstile> f)" |
    35 "s \<Turnstile> EF f    = (\<exists>t. (s,t) \<in> M\<^sup>* \<and> t \<Turnstile> f)"
    36 
    37 text{*\noindent
    38 The first three equations should be self-explanatory. The temporal formula
    39 @{term"AX f"} means that @{term f} is true in \emph{A}ll ne\emph{X}t states whereas
    40 @{term"EF f"} means that there \emph{E}xists some \emph{F}uture state in which @{term f} is
    41 true. The future is expressed via @{text"\<^sup>*"}, the reflexive transitive
    42 closure. Because of reflexivity, the future includes the present.
    43 
    44 Now we come to the model checker itself. It maps a formula into the
    45 set of states where the formula is true.  It too is defined by
    46 recursion over the syntax: *}
    47 
    48 primrec mc :: "formula \<Rightarrow> state set" where
    49 "mc(Atom a)  = {s. a \<in> L s}" |
    50 "mc(Neg f)   = -mc f" |
    51 "mc(And f g) = mc f \<inter> mc g" |
    52 "mc(AX f)    = {s. \<forall>t. (s,t) \<in> M  \<longrightarrow> t \<in> mc f}" |
    53 "mc(EF f)    = lfp(\<lambda>T. mc f \<union> (M\<inverse> `` T))"
    54 
    55 text{*\noindent
    56 Only the equation for @{term EF} deserves some comments. Remember that the
    57 postfix @{text"\<inverse>"} and the infix @{text"``"} are predefined and denote the
    58 converse of a relation and the image of a set under a relation.  Thus
    59 @{term "M\<inverse> `` T"} is the set of all predecessors of @{term T} and the least
    60 fixed point (@{term lfp}) of @{term"\<lambda>T. mc f \<union> M\<inverse> `` T"} is the least set
    61 @{term T} containing @{term"mc f"} and all predecessors of @{term T}. If you
    62 find it hard to see that @{term"mc(EF f)"} contains exactly those states from
    63 which there is a path to a state where @{term f} is true, do not worry --- this
    64 will be proved in a moment.
    65 
    66 First we prove monotonicity of the function inside @{term lfp}
    67 in order to make sure it really has a least fixed point.
    68 *}
    69 
    70 lemma mono_ef: "mono(\<lambda>T. A \<union> (M\<inverse> `` T))"
    71 apply(rule monoI)
    72 apply blast
    73 done
    74 
    75 text{*\noindent
    76 Now we can relate model checking and semantics. For the @{text EF} case we need
    77 a separate lemma:
    78 *}
    79 
    80 lemma EF_lemma:
    81   "lfp(\<lambda>T. A \<union> (M\<inverse> `` T)) = {s. \<exists>t. (s,t) \<in> M\<^sup>* \<and> t \<in> A}"
    82 
    83 txt{*\noindent
    84 The equality is proved in the canonical fashion by proving that each set
    85 includes the other; the inclusion is shown pointwise:
    86 *}
    87 
    88 apply(rule equalityI)
    89  apply(rule subsetI)
    90  apply(simp)(*<*)apply(rename_tac s)(*>*)
    91 
    92 txt{*\noindent
    93 Simplification leaves us with the following first subgoal
    94 @{subgoals[display,indent=0,goals_limit=1]}
    95 which is proved by @{term lfp}-induction:
    96 *}
    97 
    98  apply(erule lfp_induct_set)
    99   apply(rule mono_ef)
   100  apply(simp)
   101 (*pr(latex xsymbols symbols);*)
   102 txt{*\noindent
   103 Having disposed of the monotonicity subgoal,
   104 simplification leaves us with the following goal:
   105 \begin{isabelle}
   106 \ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ x\ {\isasymin}\ A\ {\isasymor}\isanewline
   107 \ \ \ \ \ \ \ \ \ x\ {\isasymin}\ M{\isasyminverse}\ {\isacharbackquote}{\isacharbackquote}\ {\isacharparenleft}lfp\ {\isacharparenleft}\dots{\isacharparenright}\ {\isasyminter}\ {\isacharbraceleft}x{\isachardot}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}x{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A{\isacharbraceright}{\isacharparenright}\isanewline
   108 \ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}x{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A
   109 \end{isabelle}
   110 It is proved by @{text blast}, using the transitivity of 
   111 \isa{M\isactrlsup {\isacharasterisk}}.
   112 *}
   113 
   114  apply(blast intro: rtrancl_trans)
   115 
   116 txt{*
   117 We now return to the second set inclusion subgoal, which is again proved
   118 pointwise:
   119 *}
   120 
   121 apply(rule subsetI)
   122 apply(simp, clarify)
   123 
   124 txt{*\noindent
   125 After simplification and clarification we are left with
   126 @{subgoals[display,indent=0,goals_limit=1]}
   127 This goal is proved by induction on @{term"(s,t)\<in>M\<^sup>*"}. But since the model
   128 checker works backwards (from @{term t} to @{term s}), we cannot use the
   129 induction theorem @{thm[source]rtrancl_induct}: it works in the
   130 forward direction. Fortunately the converse induction theorem
   131 @{thm[source]converse_rtrancl_induct} already exists:
   132 @{thm[display,margin=60]converse_rtrancl_induct[no_vars]}
   133 It says that if @{prop"(a,b):r\<^sup>*"} and we know @{prop"P b"} then we can infer
   134 @{prop"P a"} provided each step backwards from a predecessor @{term z} of
   135 @{term b} preserves @{term P}.
   136 *}
   137 
   138 apply(erule converse_rtrancl_induct)
   139 
   140 txt{*\noindent
   141 The base case
   142 @{subgoals[display,indent=0,goals_limit=1]}
   143 is solved by unrolling @{term lfp} once
   144 *}
   145 
   146  apply(subst lfp_unfold[OF mono_ef])
   147 
   148 txt{*
   149 @{subgoals[display,indent=0,goals_limit=1]}
   150 and disposing of the resulting trivial subgoal automatically:
   151 *}
   152 
   153  apply(blast)
   154 
   155 txt{*\noindent
   156 The proof of the induction step is identical to the one for the base case:
   157 *}
   158 
   159 apply(subst lfp_unfold[OF mono_ef])
   160 apply(blast)
   161 done
   162 
   163 text{*
   164 The main theorem is proved in the familiar manner: induction followed by
   165 @{text auto} augmented with the lemma as a simplification rule.
   166 *}
   167 
   168 theorem "mc f = {s. s \<Turnstile> f}"
   169 apply(induct_tac f)
   170 apply(auto simp add: EF_lemma)
   171 done
   172 
   173 text{*
   174 \begin{exercise}
   175 @{term AX} has a dual operator @{term EN} 
   176 (``there exists a next state such that'')%
   177 \footnote{We cannot use the customary @{text EX}: it is reserved
   178 as the \textsc{ascii}-equivalent of @{text"\<exists>"}.}
   179 with the intended semantics
   180 @{prop[display]"(s \<Turnstile> EN f) = (EX t. (s,t) : M & t \<Turnstile> f)"}
   181 Fortunately, @{term"EN f"} can already be expressed as a PDL formula. How?
   182 
   183 Show that the semantics for @{term EF} satisfies the following recursion equation:
   184 @{prop[display]"(s \<Turnstile> EF f) = (s \<Turnstile> f | s \<Turnstile> EN(EF f))"}
   185 \end{exercise}
   186 \index{PDL|)}
   187 *}
   188 (*<*)
   189 theorem main: "mc f = {s. s \<Turnstile> f}"
   190 apply(induct_tac f)
   191 apply(auto simp add: EF_lemma)
   192 done
   193 
   194 lemma aux: "s \<Turnstile> f = (s : mc f)"
   195 apply(simp add: main)
   196 done
   197 
   198 lemma "(s \<Turnstile> EF f) = (s \<Turnstile> f | s \<Turnstile> Neg(AX(Neg(EF f))))"
   199 apply(simp only: aux)
   200 apply(simp)
   201 apply(subst lfp_unfold[OF mono_ef], fast)
   202 done
   203 
   204 end
   205 (*>*)