10159
|
1 |
(*<*)theory CTL = Base:;(*>*)
|
9958
|
2 |
|
10971
|
3 |
subsection{*Computation Tree Logic --- CTL*};
|
9958
|
4 |
|
10212
|
5 |
text{*\label{sec:CTL}
|
11494
|
6 |
\index{CTL|(}%
|
10867
|
7 |
The semantics of PDL only needs reflexive transitive closure.
|
|
8 |
Let us be adventurous and introduce a more expressive temporal operator.
|
|
9 |
We extend the datatype
|
10149
|
10 |
@{text formula} by a new constructor
|
10159
|
11 |
*};
|
10149
|
12 |
(*<*)
|
|
13 |
datatype formula = Atom atom
|
|
14 |
| Neg formula
|
|
15 |
| And formula formula
|
|
16 |
| AX formula
|
|
17 |
| EF formula(*>*)
|
10159
|
18 |
| AF formula;
|
9958
|
19 |
|
10149
|
20 |
text{*\noindent
|
10983
|
21 |
which stands for ``\emph{A}lways in the \emph{F}uture'':
|
|
22 |
on all infinite paths, at some point the formula holds.
|
|
23 |
Formalizing the notion of an infinite path is easy
|
10159
|
24 |
in HOL: it is simply a function from @{typ nat} to @{typ state}.
|
|
25 |
*};
|
9958
|
26 |
|
|
27 |
constdefs Paths :: "state \<Rightarrow> (nat \<Rightarrow> state)set"
|
10149
|
28 |
"Paths s \<equiv> {p. s = p 0 \<and> (\<forall>i. (p i, p(i+1)) \<in> M)}";
|
|
29 |
|
|
30 |
text{*\noindent
|
11494
|
31 |
This definition allows a succinct statement of the semantics of @{term AF}:
|
10867
|
32 |
\footnote{Do not be misled: neither datatypes nor recursive functions can be
|
10149
|
33 |
extended by new constructors or equations. This is just a trick of the
|
|
34 |
presentation. In reality one has to define a new datatype and a new function.}
|
10159
|
35 |
*};
|
10149
|
36 |
(*<*)
|
10159
|
37 |
consts valid :: "state \<Rightarrow> formula \<Rightarrow> bool" ("(_ \<Turnstile> _)" [80,80] 80);
|
9958
|
38 |
|
|
39 |
primrec
|
10133
|
40 |
"s \<Turnstile> Atom a = (a \<in> L s)"
|
10149
|
41 |
"s \<Turnstile> Neg f = (~(s \<Turnstile> f))"
|
9958
|
42 |
"s \<Turnstile> And f g = (s \<Turnstile> f \<and> s \<Turnstile> g)"
|
|
43 |
"s \<Turnstile> AX f = (\<forall>t. (s,t) \<in> M \<longrightarrow> t \<Turnstile> f)"
|
10801
|
44 |
"s \<Turnstile> EF f = (\<exists>t. (s,t) \<in> M\<^sup>* \<and> t \<Turnstile> f)"
|
10149
|
45 |
(*>*)
|
9958
|
46 |
"s \<Turnstile> AF f = (\<forall>p \<in> Paths s. \<exists>i. p i \<Turnstile> f)";
|
|
47 |
|
10149
|
48 |
text{*\noindent
|
|
49 |
Model checking @{term AF} involves a function which
|
10159
|
50 |
is just complicated enough to warrant a separate definition:
|
|
51 |
*};
|
10149
|
52 |
|
9958
|
53 |
constdefs af :: "state set \<Rightarrow> state set \<Rightarrow> state set"
|
10149
|
54 |
"af A T \<equiv> A \<union> {s. \<forall>t. (s, t) \<in> M \<longrightarrow> t \<in> T}";
|
|
55 |
|
|
56 |
text{*\noindent
|
10867
|
57 |
Now we define @{term "mc(AF f)"} as the least set @{term T} that includes
|
10149
|
58 |
@{term"mc f"} and all states all of whose direct successors are in @{term T}:
|
10159
|
59 |
*};
|
10149
|
60 |
(*<*)
|
|
61 |
consts mc :: "formula \<Rightarrow> state set";
|
9958
|
62 |
primrec
|
10133
|
63 |
"mc(Atom a) = {s. a \<in> L s}"
|
10149
|
64 |
"mc(Neg f) = -mc f"
|
9958
|
65 |
"mc(And f g) = mc f \<inter> mc g"
|
|
66 |
"mc(AX f) = {s. \<forall>t. (s,t) \<in> M \<longrightarrow> t \<in> mc f}"
|
10839
|
67 |
"mc(EF f) = lfp(\<lambda>T. mc f \<union> M\<inverse> `` T)"(*>*)
|
9958
|
68 |
"mc(AF f) = lfp(af(mc f))";
|
10159
|
69 |
|
|
70 |
text{*\noindent
|
|
71 |
Because @{term af} is monotone in its second argument (and also its first, but
|
10983
|
72 |
that is irrelevant), @{term"af A"} has a least fixed point:
|
10159
|
73 |
*};
|
|
74 |
|
|
75 |
lemma mono_af: "mono(af A)";
|
|
76 |
apply(simp add: mono_def af_def);
|
|
77 |
apply blast;
|
|
78 |
done
|
10149
|
79 |
(*<*)
|
10839
|
80 |
lemma mono_ef: "mono(\<lambda>T. A \<union> M\<inverse> `` T)";
|
10159
|
81 |
apply(rule monoI);
|
|
82 |
by(blast);
|
9958
|
83 |
|
10149
|
84 |
lemma EF_lemma:
|
10839
|
85 |
"lfp(\<lambda>T. A \<union> M\<inverse> `` T) = {s. \<exists>t. (s,t) \<in> M\<^sup>* \<and> t \<in> A}";
|
9958
|
86 |
apply(rule equalityI);
|
|
87 |
apply(rule subsetI);
|
10159
|
88 |
apply(simp);
|
10210
|
89 |
apply(erule lfp_induct);
|
10159
|
90 |
apply(rule mono_ef);
|
|
91 |
apply(simp);
|
10281
|
92 |
apply(blast intro: rtrancl_trans);
|
10159
|
93 |
apply(rule subsetI);
|
|
94 |
apply(simp, clarify);
|
|
95 |
apply(erule converse_rtrancl_induct);
|
11231
|
96 |
apply(subst lfp_unfold[OF mono_ef]);
|
10159
|
97 |
apply(blast);
|
11231
|
98 |
apply(subst lfp_unfold[OF mono_ef]);
|
10159
|
99 |
by(blast);
|
10149
|
100 |
(*>*)
|
10159
|
101 |
text{*
|
10867
|
102 |
All we need to prove now is @{prop"mc(AF f) = {s. s \<Turnstile> AF f}"}, which states
|
|
103 |
that @{term mc} and @{text"\<Turnstile>"} agree for @{term AF}\@.
|
|
104 |
This time we prove the two inclusions separately, starting
|
10159
|
105 |
with the easy one:
|
|
106 |
*};
|
9958
|
107 |
|
10159
|
108 |
theorem AF_lemma1:
|
12328
|
109 |
"lfp(af A) \<subseteq> {s. \<forall>p \<in> Paths s. \<exists>i. p i \<in> A}";
|
10159
|
110 |
|
|
111 |
txt{*\noindent
|
11494
|
112 |
In contrast to the analogous proof for @{term EF}, and just
|
|
113 |
for a change, we do not use fixed point induction. Park-induction,
|
|
114 |
named after David Park, is weaker but sufficient for this proof:
|
10995
|
115 |
\begin{center}
|
|
116 |
@{thm lfp_lowerbound[of _ "S",no_vars]} \hfill (@{thm[source]lfp_lowerbound})
|
|
117 |
\end{center}
|
10225
|
118 |
The instance of the premise @{prop"f S \<subseteq> S"} is proved pointwise,
|
10281
|
119 |
a decision that clarification takes for us:
|
10159
|
120 |
*};
|
10225
|
121 |
apply(rule lfp_lowerbound);
|
10159
|
122 |
apply(clarsimp simp add: af_def Paths_def);
|
10363
|
123 |
|
10225
|
124 |
txt{*
|
10363
|
125 |
@{subgoals[display,indent=0,margin=70,goals_limit=1]}
|
10159
|
126 |
Now we eliminate the disjunction. The case @{prop"p 0 \<in> A"} is trivial:
|
|
127 |
*};
|
|
128 |
|
9958
|
129 |
apply(erule disjE);
|
|
130 |
apply(blast);
|
10159
|
131 |
|
|
132 |
txt{*\noindent
|
|
133 |
In the other case we set @{term t} to @{term"p 1"} and simplify matters:
|
|
134 |
*};
|
|
135 |
|
9958
|
136 |
apply(erule_tac x = "p 1" in allE);
|
|
137 |
apply(clarsimp);
|
10159
|
138 |
|
|
139 |
txt{*
|
10363
|
140 |
@{subgoals[display,indent=0,margin=70,goals_limit=1]}
|
11494
|
141 |
It merely remains to set @{term pa} to @{term"\<lambda>i. p(i+1)"}, that is,
|
|
142 |
@{term p} without its first element. The rest is automatic:
|
10159
|
143 |
*};
|
|
144 |
|
9958
|
145 |
apply(erule_tac x = "\<lambda>i. p(i+1)" in allE);
|
11494
|
146 |
apply force;
|
10159
|
147 |
done;
|
9958
|
148 |
|
10225
|
149 |
|
9958
|
150 |
text{*
|
10867
|
151 |
The opposite inclusion is proved by contradiction: if some state
|
10159
|
152 |
@{term s} is not in @{term"lfp(af A)"}, then we can construct an
|
11494
|
153 |
infinite @{term A}-avoiding path starting from~@{term s}. The reason is
|
10159
|
154 |
that by unfolding @{term lfp} we find that if @{term s} is not in
|
9958
|
155 |
@{term"lfp(af A)"}, then @{term s} is not in @{term A} and there is a
|
10983
|
156 |
direct successor of @{term s} that is again not in \mbox{@{term"lfp(af
|
|
157 |
A)"}}. Iterating this argument yields the promised infinite
|
9958
|
158 |
@{term A}-avoiding path. Let us formalize this sketch.
|
|
159 |
|
10867
|
160 |
The one-step argument in the sketch above
|
|
161 |
is proved by a variant of contraposition:
|
9958
|
162 |
*};
|
|
163 |
|
|
164 |
lemma not_in_lfp_afD:
|
10983
|
165 |
"s \<notin> lfp(af A) \<Longrightarrow> s \<notin> A \<and> (\<exists> t. (s,t) \<in> M \<and> t \<notin> lfp(af A))";
|
10235
|
166 |
apply(erule contrapos_np);
|
11231
|
167 |
apply(subst lfp_unfold[OF mono_af]);
|
10159
|
168 |
apply(simp add:af_def);
|
|
169 |
done;
|
9958
|
170 |
|
|
171 |
text{*\noindent
|
10867
|
172 |
We assume the negation of the conclusion and prove @{term"s : lfp(af A)"}.
|
10237
|
173 |
Unfolding @{term lfp} once and
|
9958
|
174 |
simplifying with the definition of @{term af} finishes the proof.
|
|
175 |
|
|
176 |
Now we iterate this process. The following construction of the desired
|
10895
|
177 |
path is parameterized by a predicate @{term Q} that should hold along the path:
|
9958
|
178 |
*};
|
|
179 |
|
|
180 |
consts path :: "state \<Rightarrow> (state \<Rightarrow> bool) \<Rightarrow> (nat \<Rightarrow> state)";
|
|
181 |
primrec
|
10895
|
182 |
"path s Q 0 = s"
|
|
183 |
"path s Q (Suc n) = (SOME t. (path s Q n,t) \<in> M \<and> Q t)";
|
9958
|
184 |
|
|
185 |
text{*\noindent
|
|
186 |
Element @{term"n+1"} on this path is some arbitrary successor
|
10895
|
187 |
@{term t} of element @{term n} such that @{term"Q t"} holds. Remember that @{text"SOME t. R t"}
|
10654
|
188 |
is some arbitrary but fixed @{term t} such that @{prop"R t"} holds (see \S\ref{sec:SOME}). Of
|
10867
|
189 |
course, such a @{term t} need not exist, but that is of no
|
|
190 |
concern to us since we will only use @{term path} when a
|
10159
|
191 |
suitable @{term t} does exist.
|
9958
|
192 |
|
10895
|
193 |
Let us show that if each state @{term s} that satisfies @{term Q}
|
|
194 |
has a successor that again satisfies @{term Q}, then there exists an infinite @{term Q}-path:
|
9958
|
195 |
*};
|
|
196 |
|
10159
|
197 |
lemma infinity_lemma:
|
10895
|
198 |
"\<lbrakk> Q s; \<forall>s. Q s \<longrightarrow> (\<exists> t. (s,t) \<in> M \<and> Q t) \<rbrakk> \<Longrightarrow>
|
|
199 |
\<exists>p\<in>Paths s. \<forall>i. Q(p i)";
|
9958
|
200 |
|
|
201 |
txt{*\noindent
|
10983
|
202 |
First we rephrase the conclusion slightly because we need to prove simultaneously
|
|
203 |
both the path property and the fact that @{term Q} holds:
|
9958
|
204 |
*};
|
|
205 |
|
11705
|
206 |
apply(subgoal_tac "\<exists>p. s = p (0::nat) \<and> (\<forall>i. (p i, p(i+1)) \<in> M \<and> Q(p i))");
|
9958
|
207 |
|
|
208 |
txt{*\noindent
|
10159
|
209 |
From this proposition the original goal follows easily:
|
9958
|
210 |
*};
|
|
211 |
|
|
212 |
apply(simp add:Paths_def, blast);
|
10159
|
213 |
|
|
214 |
txt{*\noindent
|
10895
|
215 |
The new subgoal is proved by providing the witness @{term "path s Q"} for @{term p}:
|
10159
|
216 |
*};
|
|
217 |
|
10895
|
218 |
apply(rule_tac x = "path s Q" in exI);
|
10159
|
219 |
apply(clarsimp);
|
|
220 |
|
|
221 |
txt{*\noindent
|
11494
|
222 |
After simplification and clarification, the subgoal has the following form:
|
10363
|
223 |
@{subgoals[display,indent=0,margin=70,goals_limit=1]}
|
11494
|
224 |
It invites a proof by induction on @{term i}:
|
10159
|
225 |
*};
|
|
226 |
|
9958
|
227 |
apply(induct_tac i);
|
|
228 |
apply(simp);
|
10159
|
229 |
|
|
230 |
txt{*\noindent
|
10983
|
231 |
After simplification, the base case boils down to
|
10363
|
232 |
@{subgoals[display,indent=0,margin=70,goals_limit=1]}
|
10159
|
233 |
The conclusion looks exceedingly trivial: after all, @{term t} is chosen such that @{prop"(s,t):M"}
|
|
234 |
holds. However, we first have to show that such a @{term t} actually exists! This reasoning
|
|
235 |
is embodied in the theorem @{thm[source]someI2_ex}:
|
|
236 |
@{thm[display,eta_contract=false]someI2_ex}
|
|
237 |
When we apply this theorem as an introduction rule, @{text"?P x"} becomes
|
10895
|
238 |
@{prop"(s, x) : M & Q x"} and @{text"?Q x"} becomes @{prop"(s,x) : M"} and we have to prove
|
|
239 |
two subgoals: @{prop"EX a. (s, a) : M & Q a"}, which follows from the assumptions, and
|
|
240 |
@{prop"(s, x) : M & Q x ==> (s,x) : M"}, which is trivial. Thus it is not surprising that
|
10159
|
241 |
@{text fast} can prove the base case quickly:
|
|
242 |
*};
|
|
243 |
|
10000
|
244 |
apply(fast intro:someI2_ex);
|
10159
|
245 |
|
|
246 |
txt{*\noindent
|
11494
|
247 |
What is worth noting here is that we have used \methdx{fast} rather than
|
10212
|
248 |
@{text blast}. The reason is that @{text blast} would fail because it cannot
|
|
249 |
cope with @{thm[source]someI2_ex}: unifying its conclusion with the current
|
11149
|
250 |
subgoal is non-trivial because of the nested schematic variables. For
|
10212
|
251 |
efficiency reasons @{text blast} does not even attempt such unifications.
|
|
252 |
Although @{text fast} can in principle cope with complicated unification
|
|
253 |
problems, in practice the number of unifiers arising is often prohibitive and
|
|
254 |
the offending rule may need to be applied explicitly rather than
|
|
255 |
automatically. This is what happens in the step case.
|
10159
|
256 |
|
10212
|
257 |
The induction step is similar, but more involved, because now we face nested
|
|
258 |
occurrences of @{text SOME}. As a result, @{text fast} is no longer able to
|
|
259 |
solve the subgoal and we apply @{thm[source]someI2_ex} by hand. We merely
|
|
260 |
show the proof commands but do not describe the details:
|
10159
|
261 |
*};
|
|
262 |
|
9958
|
263 |
apply(simp);
|
10000
|
264 |
apply(rule someI2_ex);
|
9958
|
265 |
apply(blast);
|
10000
|
266 |
apply(rule someI2_ex);
|
9958
|
267 |
apply(blast);
|
10159
|
268 |
apply(blast);
|
|
269 |
done;
|
9958
|
270 |
|
10159
|
271 |
text{*
|
10867
|
272 |
Function @{term path} has fulfilled its purpose now and can be forgotten.
|
|
273 |
It was merely defined to provide the witness in the proof of the
|
10171
|
274 |
@{thm[source]infinity_lemma}. Aficionados of minimal proofs might like to know
|
10159
|
275 |
that we could have given the witness without having to define a new function:
|
|
276 |
the term
|
10895
|
277 |
@{term[display]"nat_rec s (\<lambda>n t. SOME u. (t,u)\<in>M \<and> Q u)"}
|
|
278 |
is extensionally equal to @{term"path s Q"},
|
10867
|
279 |
where @{term nat_rec} is the predefined primitive recursor on @{typ nat}.
|
10159
|
280 |
*};
|
|
281 |
(*<*)
|
|
282 |
lemma infinity_lemma:
|
10895
|
283 |
"\<lbrakk> Q s; \<forall> s. Q s \<longrightarrow> (\<exists> t. (s,t)\<in>M \<and> Q t) \<rbrakk> \<Longrightarrow>
|
|
284 |
\<exists> p\<in>Paths s. \<forall> i. Q(p i)";
|
9958
|
285 |
apply(subgoal_tac
|
10895
|
286 |
"\<exists> p. s = p 0 \<and> (\<forall> i. (p i,p(Suc i))\<in>M \<and> Q(p i))");
|
9958
|
287 |
apply(simp add:Paths_def);
|
|
288 |
apply(blast);
|
10895
|
289 |
apply(rule_tac x = "nat_rec s (\<lambda>n t. SOME u. (t,u)\<in>M \<and> Q u)" in exI);
|
9958
|
290 |
apply(simp);
|
|
291 |
apply(intro strip);
|
|
292 |
apply(induct_tac i);
|
|
293 |
apply(simp);
|
10000
|
294 |
apply(fast intro:someI2_ex);
|
9958
|
295 |
apply(simp);
|
10000
|
296 |
apply(rule someI2_ex);
|
9958
|
297 |
apply(blast);
|
10000
|
298 |
apply(rule someI2_ex);
|
9958
|
299 |
apply(blast);
|
|
300 |
by(blast);
|
10159
|
301 |
(*>*)
|
9958
|
302 |
|
10159
|
303 |
text{*
|
|
304 |
At last we can prove the opposite direction of @{thm[source]AF_lemma1}:
|
|
305 |
*};
|
|
306 |
|
12328
|
307 |
theorem AF_lemma2: "{s. \<forall>p \<in> Paths s. \<exists>i. p i \<in> A} \<subseteq> lfp(af A)";
|
10159
|
308 |
|
|
309 |
txt{*\noindent
|
10237
|
310 |
The proof is again pointwise and then by contraposition:
|
10159
|
311 |
*};
|
|
312 |
|
9958
|
313 |
apply(rule subsetI);
|
10235
|
314 |
apply(erule contrapos_pp);
|
9958
|
315 |
apply simp;
|
10159
|
316 |
|
|
317 |
txt{*
|
10363
|
318 |
@{subgoals[display,indent=0,goals_limit=1]}
|
10159
|
319 |
Applying the @{thm[source]infinity_lemma} as a destruction rule leaves two subgoals, the second
|
|
320 |
premise of @{thm[source]infinity_lemma} and the original subgoal:
|
|
321 |
*};
|
|
322 |
|
|
323 |
apply(drule infinity_lemma);
|
|
324 |
|
|
325 |
txt{*
|
10363
|
326 |
@{subgoals[display,indent=0,margin=65]}
|
10159
|
327 |
Both are solved automatically:
|
|
328 |
*};
|
9958
|
329 |
|
10159
|
330 |
apply(auto dest:not_in_lfp_afD);
|
|
331 |
done;
|
9958
|
332 |
|
10159
|
333 |
text{*
|
10867
|
334 |
If you find these proofs too complicated, we recommend that you read
|
|
335 |
\S\ref{sec:CTL-revisited}, where we show how inductive definitions lead to
|
10217
|
336 |
simpler arguments.
|
|
337 |
|
|
338 |
The main theorem is proved as for PDL, except that we also derive the
|
|
339 |
necessary equality @{text"lfp(af A) = ..."} by combining
|
|
340 |
@{thm[source]AF_lemma1} and @{thm[source]AF_lemma2} on the spot:
|
10159
|
341 |
*}
|
|
342 |
|
|
343 |
theorem "mc f = {s. s \<Turnstile> f}";
|
|
344 |
apply(induct_tac f);
|
|
345 |
apply(auto simp add: EF_lemma equalityI[OF AF_lemma1 AF_lemma2]);
|
|
346 |
done
|
|
347 |
|
|
348 |
text{*
|
10281
|
349 |
|
10867
|
350 |
The language defined above is not quite CTL\@. The latter also includes an
|
10983
|
351 |
until-operator @{term"EU f g"} with semantics ``there \emph{E}xists a path
|
11494
|
352 |
where @{term f} is true \emph{U}ntil @{term g} becomes true''. We need
|
|
353 |
an auxiliary function:
|
10281
|
354 |
*}
|
|
355 |
|
|
356 |
consts until:: "state set \<Rightarrow> state set \<Rightarrow> state \<Rightarrow> state list \<Rightarrow> bool"
|
|
357 |
primrec
|
|
358 |
"until A B s [] = (s \<in> B)"
|
|
359 |
"until A B s (t#p) = (s \<in> A \<and> (s,t) \<in> M \<and> until A B t p)"
|
|
360 |
(*<*)constdefs
|
|
361 |
eusem :: "state set \<Rightarrow> state set \<Rightarrow> state set"
|
|
362 |
"eusem A B \<equiv> {s. \<exists>p. until A B s p}"(*>*)
|
|
363 |
|
|
364 |
text{*\noindent
|
11494
|
365 |
Expressing the semantics of @{term EU} is now straightforward:
|
10983
|
366 |
@{prop[display]"s \<Turnstile> EU f g = (\<exists>p. until {t. t \<Turnstile> f} {t. t \<Turnstile> g} s p)"}
|
10281
|
367 |
Note that @{term EU} is not definable in terms of the other operators!
|
|
368 |
|
|
369 |
Model checking @{term EU} is again a least fixed point construction:
|
10839
|
370 |
@{text[display]"mc(EU f g) = lfp(\<lambda>T. mc g \<union> mc f \<inter> (M\<inverse> `` T))"}
|
10281
|
371 |
|
10171
|
372 |
\begin{exercise}
|
10281
|
373 |
Extend the datatype of formulae by the above until operator
|
|
374 |
and prove the equivalence between semantics and model checking, i.e.\ that
|
10186
|
375 |
@{prop[display]"mc(EU f g) = {s. s \<Turnstile> EU f g}"}
|
|
376 |
%For readability you may want to annotate {term EU} with its customary syntax
|
|
377 |
%{text[display]"| EU formula formula E[_ U _]"}
|
|
378 |
%which enables you to read and write {text"E[f U g]"} instead of {term"EU f g"}.
|
|
379 |
\end{exercise}
|
10867
|
380 |
For more CTL exercises see, for example, Huth and Ryan \cite{Huth-Ryan-book}.
|
10281
|
381 |
*}
|
|
382 |
|
|
383 |
(*<*)
|
|
384 |
constdefs
|
|
385 |
eufix :: "state set \<Rightarrow> state set \<Rightarrow> state set \<Rightarrow> state set"
|
10839
|
386 |
"eufix A B T \<equiv> B \<union> A \<inter> (M\<inverse> `` T)"
|
10281
|
387 |
|
|
388 |
lemma "lfp(eufix A B) \<subseteq> eusem A B"
|
|
389 |
apply(rule lfp_lowerbound)
|
|
390 |
apply(clarsimp simp add:eusem_def eufix_def);
|
|
391 |
apply(erule disjE);
|
|
392 |
apply(rule_tac x = "[]" in exI);
|
|
393 |
apply simp
|
|
394 |
apply(clarsimp);
|
|
395 |
apply(rule_tac x = "y#xc" in exI);
|
|
396 |
apply simp;
|
|
397 |
done
|
|
398 |
|
|
399 |
lemma mono_eufix: "mono(eufix A B)";
|
|
400 |
apply(simp add: mono_def eufix_def);
|
|
401 |
apply blast;
|
|
402 |
done
|
|
403 |
|
|
404 |
lemma "eusem A B \<subseteq> lfp(eufix A B)";
|
|
405 |
apply(clarsimp simp add:eusem_def);
|
|
406 |
apply(erule rev_mp);
|
|
407 |
apply(rule_tac x = x in spec);
|
|
408 |
apply(induct_tac p);
|
11231
|
409 |
apply(subst lfp_unfold[OF mono_eufix])
|
10281
|
410 |
apply(simp add:eufix_def);
|
|
411 |
apply(clarsimp);
|
11231
|
412 |
apply(subst lfp_unfold[OF mono_eufix])
|
10281
|
413 |
apply(simp add:eufix_def);
|
|
414 |
apply blast;
|
|
415 |
done
|
10178
|
416 |
|
10281
|
417 |
(*
|
|
418 |
constdefs
|
|
419 |
eusem :: "state set \<Rightarrow> state set \<Rightarrow> state set"
|
|
420 |
"eusem A B \<equiv> {s. \<exists>p\<in>Paths s. \<exists>j. p j \<in> B \<and> (\<forall>i < j. p i \<in> A)}"
|
|
421 |
|
|
422 |
axioms
|
|
423 |
M_total: "\<exists>t. (s,t) \<in> M"
|
|
424 |
|
|
425 |
consts apath :: "state \<Rightarrow> (nat \<Rightarrow> state)"
|
|
426 |
primrec
|
|
427 |
"apath s 0 = s"
|
|
428 |
"apath s (Suc i) = (SOME t. (apath s i,t) \<in> M)"
|
|
429 |
|
|
430 |
lemma [iff]: "apath s \<in> Paths s";
|
|
431 |
apply(simp add:Paths_def);
|
|
432 |
apply(blast intro: M_total[THEN someI_ex])
|
|
433 |
done
|
|
434 |
|
|
435 |
constdefs
|
|
436 |
pcons :: "state \<Rightarrow> (nat \<Rightarrow> state) \<Rightarrow> (nat \<Rightarrow> state)"
|
|
437 |
"pcons s p == \<lambda>i. case i of 0 \<Rightarrow> s | Suc j \<Rightarrow> p j"
|
|
438 |
|
|
439 |
lemma pcons_PathI: "[| (s,t) : M; p \<in> Paths t |] ==> pcons s p \<in> Paths s";
|
|
440 |
by(simp add:Paths_def pcons_def split:nat.split);
|
|
441 |
|
|
442 |
lemma "lfp(eufix A B) \<subseteq> eusem A B"
|
|
443 |
apply(rule lfp_lowerbound)
|
|
444 |
apply(clarsimp simp add:eusem_def eufix_def);
|
|
445 |
apply(erule disjE);
|
|
446 |
apply(rule_tac x = "apath x" in bexI);
|
|
447 |
apply(rule_tac x = 0 in exI);
|
|
448 |
apply simp;
|
|
449 |
apply simp;
|
|
450 |
apply(clarify);
|
|
451 |
apply(rule_tac x = "pcons xb p" in bexI);
|
|
452 |
apply(rule_tac x = "j+1" in exI);
|
|
453 |
apply (simp add:pcons_def split:nat.split);
|
|
454 |
apply (simp add:pcons_PathI)
|
|
455 |
done
|
|
456 |
*)
|
|
457 |
(*>*)
|
12334
|
458 |
|
|
459 |
text{* Let us close this section with a few words about the executability of
|
|
460 |
our model checkers. It is clear that if all sets are finite, they can be
|
|
461 |
represented as lists and the usual set operations are easily
|
|
462 |
implemented. Only @{term lfp} requires a little thought. Fortunately, theory
|
12473
|
463 |
@{text While_Combinator} in the Library~\cite{HOL-Library} provides a
|
12334
|
464 |
theorem stating that in the case of finite sets and a monotone
|
|
465 |
function~@{term F}, the value of \mbox{@{term"lfp F"}} can be computed by
|
|
466 |
iterated application of @{term F} to~@{term"{}"} until a fixed point is
|
|
467 |
reached. It is actually possible to generate executable functional programs
|
11494
|
468 |
from HOL definitions, but that is beyond the scope of the tutorial.%
|
12334
|
469 |
\index{CTL|)} *}
|
10212
|
470 |
(*<*)end(*>*)
|