| 
17914
 | 
     1  | 
(*<*)theory PDL imports Base begin(*>*)
  | 
| 
9958
 | 
     2  | 
  | 
| 
10971
 | 
     3  | 
subsection{*Propositional Dynamic Logic --- PDL*}
 | 
| 
9958
 | 
     4  | 
  | 
| 
10178
 | 
     5  | 
text{*\index{PDL|(}
 | 
| 
11458
 | 
     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
 | 
| 
11207
 | 
    11  | 
\cite{HarelKT-DL} looks quite different from ours, but the two are easily
 | 
| 
11458
 | 
    12  | 
shown to be equivalent.}
  | 
| 
10133
 | 
    13  | 
*}
  | 
| 
9958
 | 
    14  | 
  | 
| 
18724
 | 
    15  | 
datatype formula = Atom "atom"
  | 
| 
10149
 | 
    16  | 
                  | Neg formula
  | 
| 
 | 
    17  | 
                  | And formula formula
  | 
| 
 | 
    18  | 
                  | AX formula
  | 
| 
 | 
    19  | 
                  | EF formula
  | 
| 
10133
 | 
    20  | 
  | 
| 
 | 
    21  | 
text{*\noindent
 | 
| 
11458
 | 
    22  | 
This resembles the boolean expression case study in
  | 
| 
10867
 | 
    23  | 
\S\ref{sec:boolex}.
 | 
| 
27015
 | 
    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:
 | 
| 
10133
 | 
    27  | 
*}
  | 
| 
 | 
    28  | 
  | 
| 
27015
 | 
    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)" |
  | 
| 
12631
 | 
    35  | 
"s \<Turnstile> EF f    = (\<exists>t. (s,t) \<in> M\<^sup>* \<and> t \<Turnstile> f)"
  | 
| 
9958
 | 
    36  | 
  | 
| 
10149
 | 
    37  | 
text{*\noindent
 | 
| 
 | 
    38  | 
The first three equations should be self-explanatory. The temporal formula
  | 
| 
10983
 | 
    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
 | 
| 
10867
 | 
    41  | 
true. The future is expressed via @{text"\<^sup>*"}, the reflexive transitive
 | 
| 
10149
 | 
    42  | 
closure. Because of reflexivity, the future includes the present.
  | 
| 
 | 
    43  | 
  | 
| 
27015
 | 
    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: *}
  | 
| 
10133
 | 
    47  | 
  | 
| 
27015
 | 
    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}" |
 | 
| 
10867
 | 
    53  | 
"mc(EF f)    = lfp(\<lambda>T. mc f \<union> (M\<inverse> `` T))"
  | 
| 
10133
 | 
    54  | 
  | 
| 
10149
 | 
    55  | 
text{*\noindent
 | 
| 
 | 
    56  | 
Only the equation for @{term EF} deserves some comments. Remember that the
 | 
| 
10839
 | 
    57  | 
postfix @{text"\<inverse>"} and the infix @{text"``"} are predefined and denote the
 | 
| 
10867
 | 
    58  | 
converse of a relation and the image of a set under a relation.  Thus
  | 
| 
10839
 | 
    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
 | 
| 
10149
 | 
    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
 | 
| 
10983
 | 
    63  | 
which there is a path to a state where @{term f} is true, do not worry --- this
 | 
| 
10149
 | 
    64  | 
will be proved in a moment.
  | 
| 
 | 
    65  | 
  | 
| 
 | 
    66  | 
First we prove monotonicity of the function inside @{term lfp}
 | 
| 
10867
 | 
    67  | 
in order to make sure it really has a least fixed point.
  | 
| 
10133
 | 
    68  | 
*}
  | 
| 
 | 
    69  | 
  | 
| 
10867
 | 
    70  | 
lemma mono_ef: "mono(\<lambda>T. A \<union> (M\<inverse> `` T))"
  | 
| 
10149
 | 
    71  | 
apply(rule monoI)
  | 
| 
10159
 | 
    72  | 
apply blast
  | 
| 
 | 
    73  | 
done
  | 
| 
10149
 | 
    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:
  | 
| 
10867
 | 
    81  | 
  "lfp(\<lambda>T. A \<union> (M\<inverse> `` T)) = {s. \<exists>t. (s,t) \<in> M\<^sup>* \<and> t \<in> A}"
 | 
| 
10149
 | 
    82  | 
  | 
| 
 | 
    83  | 
txt{*\noindent
 | 
| 
 | 
    84  | 
The equality is proved in the canonical fashion by proving that each set
  | 
| 
10867
 | 
    85  | 
includes the other; the inclusion is shown pointwise:
  | 
| 
10149
 | 
    86  | 
*}
  | 
| 
 | 
    87  | 
  | 
| 
12631
 | 
    88  | 
apply(rule equalityI)
  | 
| 
 | 
    89  | 
 apply(rule subsetI)
  | 
| 
10524
 | 
    90  | 
 apply(simp)(*<*)apply(rename_tac s)(*>*)
  | 
| 
10363
 | 
    91  | 
  | 
| 
10149
 | 
    92  | 
txt{*\noindent
 | 
| 
 | 
    93  | 
Simplification leaves us with the following first subgoal
  | 
| 
10363
 | 
    94  | 
@{subgoals[display,indent=0,goals_limit=1]}
 | 
| 
10149
 | 
    95  | 
which is proved by @{term lfp}-induction:
 | 
| 
 | 
    96  | 
*}
  | 
| 
 | 
    97  | 
  | 
| 
21202
 | 
    98  | 
 apply(erule lfp_induct_set)
  | 
| 
10149
 | 
    99  | 
  apply(rule mono_ef)
  | 
| 
 | 
   100  | 
 apply(simp)
  | 
| 
 | 
   101  | 
(*pr(latex xsymbols symbols);*)
  | 
| 
 | 
   102  | 
txt{*\noindent
 | 
| 
 | 
   103  | 
Having disposed of the monotonicity subgoal,
  | 
| 
11458
 | 
   104  | 
simplification leaves us with the following goal:
  | 
| 
10149
 | 
   105  | 
\begin{isabelle}
 | 
| 
10801
 | 
   106  | 
\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x{\isachardot}\ x\ {\isasymin}\ A\ {\isasymor}\isanewline
 | 
| 
10895
 | 
   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
 | 
| 
10801
 | 
   108  | 
\ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ {\isasymexists}t{\isachardot}\ {\isacharparenleft}x{\isacharcomma}\ t{\isacharparenright}\ {\isasymin}\ M\isactrlsup {\isacharasterisk}\ {\isasymand}\ t\ {\isasymin}\ A
 | 
| 
10149
 | 
   109  | 
\end{isabelle}
 | 
| 
11458
 | 
   110  | 
It is proved by @{text blast}, using the transitivity of 
 | 
| 
 | 
   111  | 
\isa{M\isactrlsup {\isacharasterisk}}.
 | 
| 
10149
 | 
   112  | 
*}
  | 
| 
 | 
   113  | 
  | 
| 
12631
 | 
   114  | 
 apply(blast intro: rtrancl_trans)
  | 
| 
10149
 | 
   115  | 
  | 
| 
 | 
   116  | 
txt{*
 | 
| 
10867
 | 
   117  | 
We now return to the second set inclusion subgoal, which is again proved
  | 
| 
10149
 | 
   118  | 
pointwise:
  | 
| 
 | 
   119  | 
*}
  | 
| 
 | 
   120  | 
  | 
| 
 | 
   121  | 
apply(rule subsetI)
  | 
| 
 | 
   122  | 
apply(simp, clarify)
  | 
| 
10363
 | 
   123  | 
  | 
| 
10149
 | 
   124  | 
txt{*\noindent
 | 
| 
 | 
   125  | 
After simplification and clarification we are left with
  | 
| 
10363
 | 
   126  | 
@{subgoals[display,indent=0,goals_limit=1]}
 | 
| 
10801
 | 
   127  | 
This goal is proved by induction on @{term"(s,t)\<in>M\<^sup>*"}. But since the model
 | 
| 
10149
 | 
   128  | 
checker works backwards (from @{term t} to @{term s}), we cannot use the
 | 
| 
11458
 | 
   129  | 
induction theorem @{thm[source]rtrancl_induct}: it works in the
 | 
| 
10149
 | 
   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]}
 | 
| 
10801
 | 
   133  | 
It says that if @{prop"(a,b):r\<^sup>*"} and we know @{prop"P b"} then we can infer
 | 
| 
10149
 | 
   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)
  | 
| 
10363
 | 
   139  | 
  | 
| 
10149
 | 
   140  | 
txt{*\noindent
 | 
| 
 | 
   141  | 
The base case
  | 
| 
10363
 | 
   142  | 
@{subgoals[display,indent=0,goals_limit=1]}
 | 
| 
10149
 | 
   143  | 
is solved by unrolling @{term lfp} once
 | 
| 
 | 
   144  | 
*}
  | 
| 
 | 
   145  | 
  | 
| 
11231
 | 
   146  | 
 apply(subst lfp_unfold[OF mono_ef])
  | 
| 
10363
 | 
   147  | 
  | 
| 
10149
 | 
   148  | 
txt{*
 | 
| 
10363
 | 
   149  | 
@{subgoals[display,indent=0,goals_limit=1]}
 | 
| 
10149
 | 
   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  | 
  | 
| 
11231
 | 
   159  | 
apply(subst lfp_unfold[OF mono_ef])
  | 
| 
10159
 | 
   160  | 
apply(blast)
  | 
| 
 | 
   161  | 
done
  | 
| 
10149
 | 
   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  | 
*}
  | 
| 
9958
 | 
   167  | 
  | 
| 
12631
 | 
   168  | 
theorem "mc f = {s. s \<Turnstile> f}"
 | 
| 
 | 
   169  | 
apply(induct_tac f)
  | 
| 
12815
 | 
   170  | 
apply(auto simp add: EF_lemma)
  | 
| 
12631
 | 
   171  | 
done
  | 
| 
10171
 | 
   172  | 
  | 
| 
 | 
   173  | 
text{*
 | 
| 
 | 
   174  | 
\begin{exercise}
 | 
| 
11458
 | 
   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
  | 
| 
10171
 | 
   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}
 | 
| 
10178
 | 
   186  | 
\index{PDL|)}
 | 
| 
10171
 | 
   187  | 
*}
  | 
| 
 | 
   188  | 
(*<*)
  | 
| 
12631
 | 
   189  | 
theorem main: "mc f = {s. s \<Turnstile> f}"
 | 
| 
 | 
   190  | 
apply(induct_tac f)
  | 
| 
 | 
   191  | 
apply(auto simp add: EF_lemma)
  | 
| 
 | 
   192  | 
done
  | 
| 
10171
 | 
   193  | 
  | 
| 
12631
 | 
   194  | 
lemma aux: "s \<Turnstile> f = (s : mc f)"
  | 
| 
 | 
   195  | 
apply(simp add: main)
  | 
| 
 | 
   196  | 
done
  | 
| 
10171
 | 
   197  | 
  | 
| 
12631
 | 
   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)
  | 
| 
10171
 | 
   202  | 
done
  | 
| 
 | 
   203  | 
  | 
| 
 | 
   204  | 
end
  | 
| 
 | 
   205  | 
(*>*)
  |