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