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