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