262
|
1 |
(*<*)
|
|
2 |
theory PrioGDef
|
|
3 |
imports Precedence_ord Moment
|
|
4 |
begin
|
|
5 |
(*>*)
|
|
6 |
|
|
7 |
text {*
|
|
8 |
In this section, the formal model of Priority Inheritance is presented.
|
|
9 |
The model is based on Paulson's inductive protocol verification method, where
|
|
10 |
the state of the system is modelled as a list of events happened so far with the latest
|
|
11 |
event put at the head.
|
|
12 |
|
|
13 |
To define events, the identifiers of {\em threads},
|
|
14 |
{\em priority} and {\em critical resources } (abbreviated as @{text "cs"})
|
|
15 |
need to be represented. All three are represetned using standard
|
|
16 |
Isabelle/HOL type @{typ "nat"}:
|
|
17 |
*}
|
|
18 |
|
|
19 |
type_synonym thread = nat -- {* Type for thread identifiers. *}
|
|
20 |
type_synonym priority = nat -- {* Type for priorities. *}
|
|
21 |
type_synonym cs = nat -- {* Type for critical sections (or critical resources). *}
|
|
22 |
|
|
23 |
text {*
|
|
24 |
\noindent
|
|
25 |
Every event in the system corresponds to a system call, the formats of which are
|
|
26 |
defined as follows:
|
|
27 |
*}
|
|
28 |
|
|
29 |
datatype event =
|
|
30 |
Create thread priority | -- {* Thread @{text "thread"} is created with priority @{text "priority"}. *}
|
|
31 |
Exit thread | -- {* Thread @{text "thread"} finishing its execution. *}
|
|
32 |
P thread cs | -- {* Thread @{text "thread"} requesting critical resource @{text "cs"}. *}
|
|
33 |
V thread cs | -- {* Thread @{text "thread"} releasing critical resource @{text "cs"}. *}
|
|
34 |
Set thread priority -- {* Thread @{text "thread"} resets its priority to @{text "priority"}. *}
|
|
35 |
|
|
36 |
text {*
|
|
37 |
\noindent
|
|
38 |
Resource Allocation Graph (RAG for short) is used extensively in our formal analysis.
|
|
39 |
The following type @{text "node"} is used to represent nodes in RAG.
|
|
40 |
*}
|
|
41 |
datatype node =
|
|
42 |
Th "thread" | -- {* Node for thread. *}
|
|
43 |
Cs "cs" -- {* Node for critical resource. *}
|
|
44 |
|
|
45 |
text {*
|
|
46 |
In Paulson's inductive method, the states of system are represented as lists of events,
|
|
47 |
which is defined by the following type @{text "state"}:
|
|
48 |
*}
|
|
49 |
type_synonym state = "event list"
|
|
50 |
|
|
51 |
text {*
|
|
52 |
\noindent
|
|
53 |
The following function
|
|
54 |
@{text "threads"} is used to calculate the set of live threads (@{text "threads s"})
|
|
55 |
in state @{text "s"}.
|
|
56 |
*}
|
|
57 |
fun threads :: "state \<Rightarrow> thread set"
|
|
58 |
where
|
|
59 |
-- {* At the start of the system, the set of threads is empty: *}
|
|
60 |
"threads [] = {}" |
|
|
61 |
-- {* New thread is added to the @{text "threads"}: *}
|
|
62 |
"threads (Create thread prio#s) = {thread} \<union> threads s" |
|
|
63 |
-- {* Finished thread is removed: *}
|
|
64 |
"threads (Exit thread # s) = (threads s) - {thread}" |
|
|
65 |
-- {* Other kind of events does not affect the value of @{text "threads"}: *}
|
|
66 |
"threads (e#s) = threads s"
|
|
67 |
text {* \noindent
|
|
68 |
Functions such as @{text "threads"}, which extract information out of system states, are called
|
|
69 |
{\em observing functions}. A series of observing functions will be defined in the sequel in order to
|
|
70 |
model the protocol.
|
|
71 |
Observing function @{text "original_priority"} calculates
|
|
72 |
the {\em original priority} of thread @{text "th"} in state @{text "s"}, expressed as
|
|
73 |
: @{text "original_priority th s" }. The {\em original priority} is the priority
|
|
74 |
assigned to a thread when it is created or when it is reset by system call
|
|
75 |
@{text "Set thread priority"}.
|
|
76 |
*}
|
|
77 |
|
|
78 |
fun original_priority :: "thread \<Rightarrow> state \<Rightarrow> priority"
|
|
79 |
where
|
|
80 |
-- {* @{text "0"} is assigned to threads which have never been created: *}
|
|
81 |
"original_priority thread [] = 0" |
|
|
82 |
"original_priority thread (Create thread' prio#s) =
|
|
83 |
(if thread' = thread then prio else original_priority thread s)" |
|
|
84 |
"original_priority thread (Set thread' prio#s) =
|
|
85 |
(if thread' = thread then prio else original_priority thread s)" |
|
|
86 |
"original_priority thread (e#s) = original_priority thread s"
|
|
87 |
|
|
88 |
text {*
|
|
89 |
\noindent
|
|
90 |
In the following,
|
|
91 |
@{text "birthtime th s"} is the time when thread @{text "th"} is created,
|
|
92 |
observed from state @{text "s"}.
|
|
93 |
The time in the system is measured by the number of events happened so far since the very beginning.
|
|
94 |
*}
|
|
95 |
fun birthtime :: "thread \<Rightarrow> state \<Rightarrow> nat"
|
|
96 |
where
|
|
97 |
"birthtime thread [] = 0" |
|
|
98 |
"birthtime thread ((Create thread' prio)#s) =
|
|
99 |
(if (thread = thread') then length s else birthtime thread s)" |
|
|
100 |
"birthtime thread ((Set thread' prio)#s) =
|
|
101 |
(if (thread = thread') then length s else birthtime thread s)" |
|
|
102 |
"birthtime thread (e#s) = birthtime thread s"
|
|
103 |
|
|
104 |
text {*
|
|
105 |
\noindent
|
|
106 |
The {\em precedence} is a notion derived from {\em priority}, where the {\em precedence} of
|
|
107 |
a thread is the combination of its {\em original priority} and {\em birth time}. The intention is
|
|
108 |
to discriminate threads with the same priority by giving threads whose priority
|
|
109 |
is assigned earlier higher precedences, becasue such threads are more urgent to finish.
|
|
110 |
This explains the following definition:
|
|
111 |
*}
|
|
112 |
definition preced :: "thread \<Rightarrow> state \<Rightarrow> precedence"
|
|
113 |
where "preced thread s = Prc (original_priority thread s) (birthtime thread s)"
|
|
114 |
|
|
115 |
|
|
116 |
text {*
|
|
117 |
\noindent
|
|
118 |
A number of important notions are defined here:
|
|
119 |
*}
|
|
120 |
|
|
121 |
consts
|
|
122 |
holding :: "'b \<Rightarrow> thread \<Rightarrow> cs \<Rightarrow> bool"
|
|
123 |
waiting :: "'b \<Rightarrow> thread \<Rightarrow> cs \<Rightarrow> bool"
|
|
124 |
depend :: "'b \<Rightarrow> (node \<times> node) set"
|
|
125 |
dependents :: "'b \<Rightarrow> thread \<Rightarrow> thread set"
|
|
126 |
|
|
127 |
text {*
|
|
128 |
\noindent
|
|
129 |
In the definition of the following several functions, it is supposed that
|
|
130 |
the waiting queue of every critical resource is given by a waiting queue
|
|
131 |
function @{text "wq"}, which servers as arguments of these functions.
|
|
132 |
*}
|
|
133 |
defs (overloaded)
|
|
134 |
-- {*
|
|
135 |
\begin{minipage}{0.9\textwidth}
|
|
136 |
We define that the thread which is at the head of waiting queue of resource @{text "cs"}
|
|
137 |
is holding the resource. This definition is slightly different from tradition where
|
|
138 |
all threads in the waiting queue are considered as waiting for the resource.
|
|
139 |
This notion is reflected in the definition of @{text "holding wq th cs"} as follows:
|
|
140 |
\end{minipage}
|
|
141 |
*}
|
|
142 |
cs_holding_def:
|
|
143 |
"holding wq thread cs \<equiv> (thread \<in> set (wq cs) \<and> thread = hd (wq cs))"
|
|
144 |
-- {*
|
|
145 |
\begin{minipage}{0.9\textwidth}
|
|
146 |
In accordance with the definition of @{text "holding wq th cs"},
|
|
147 |
a thread @{text "th"} is considered waiting for @{text "cs"} if
|
|
148 |
it is in the {\em waiting queue} of critical resource @{text "cs"}, but not at the head.
|
|
149 |
This is reflected in the definition of @{text "waiting wq th cs"} as follows:
|
|
150 |
\end{minipage}
|
|
151 |
*}
|
|
152 |
cs_waiting_def:
|
|
153 |
"waiting wq thread cs \<equiv> (thread \<in> set (wq cs) \<and> thread \<noteq> hd (wq cs))"
|
|
154 |
-- {*
|
|
155 |
\begin{minipage}{0.9\textwidth}
|
|
156 |
@{text "depend wq"} represents the Resource Allocation Graph of the system under the waiting
|
|
157 |
queue function @{text "wq"}.
|
|
158 |
\end{minipage}
|
|
159 |
*}
|
|
160 |
cs_depend_def:
|
|
161 |
"depend (wq::cs \<Rightarrow> thread list) \<equiv>
|
|
162 |
{(Th t, Cs c) | t c. waiting wq t c} \<union> {(Cs c, Th t) | c t. holding wq t c}"
|
|
163 |
-- {*
|
|
164 |
\begin{minipage}{0.9\textwidth}
|
|
165 |
The following @{text "dependents wq th"} represents the set of threads which are depending on
|
|
166 |
thread @{text "th"} in Resource Allocation Graph @{text "depend wq"}:
|
|
167 |
\end{minipage}
|
|
168 |
*}
|
|
169 |
cs_dependents_def:
|
|
170 |
"dependents (wq::cs \<Rightarrow> thread list) th \<equiv> {th' . (Th th', Th th) \<in> (depend wq)^+}"
|
|
171 |
|
|
172 |
text {*
|
|
173 |
The data structure used by the operating system for scheduling is referred to as
|
|
174 |
{\em schedule state}. It is represented as a record consisting of
|
|
175 |
a function assigning waiting queue to resources and a function assigning precedence to
|
|
176 |
threads:
|
|
177 |
*}
|
|
178 |
record schedule_state =
|
|
179 |
waiting_queue :: "cs \<Rightarrow> thread list" -- {* The function assigning waiting queue. *}
|
|
180 |
cur_preced :: "thread \<Rightarrow> precedence" -- {* The function assigning precedence. *}
|
|
181 |
|
|
182 |
text {* \noindent
|
|
183 |
The following
|
|
184 |
@{text "cpreced s th"} gives the {\em current precedence} of thread @{text "th"} under
|
|
185 |
state @{text "s"}. The definition of @{text "cpreced"} reflects the basic idea of
|
|
186 |
Priority Inheritance that the {\em current precedence} of a thread is the precedence
|
|
187 |
inherited from the maximum of all its dependents, i.e. the threads which are waiting
|
|
188 |
directly or indirectly waiting for some resources from it. If no such thread exits,
|
|
189 |
@{text "th"}'s {\em current precedence} equals its original precedence, i.e.
|
|
190 |
@{text "preced th s"}.
|
|
191 |
*}
|
|
192 |
definition cpreced :: "state \<Rightarrow> (cs \<Rightarrow> thread list) \<Rightarrow> thread \<Rightarrow> precedence"
|
|
193 |
where "cpreced s wq = (\<lambda> th. Max ((\<lambda> th. preced th s) ` ({th} \<union> dependents wq th)))"
|
|
194 |
|
|
195 |
text {* \noindent
|
|
196 |
The following function @{text "schs"} is used to calculate the schedule state @{text "schs s"}.
|
|
197 |
It is the key function to model Priority Inheritance:
|
|
198 |
*}
|
|
199 |
fun schs :: "state \<Rightarrow> schedule_state"
|
|
200 |
where "schs [] = \<lparr>waiting_queue = \<lambda> cs. [], cur_preced = cpreced [] (\<lambda> cs. [])\<rparr>" |
|
|
201 |
-- {*
|
|
202 |
\begin{minipage}{0.9\textwidth}
|
|
203 |
\begin{enumerate}
|
|
204 |
\item @{text "ps"} is the schedule state of last moment.
|
|
205 |
\item @{text "pwq"} is the waiting queue function of last moment.
|
|
206 |
\item @{text "pcp"} is the precedence function of last moment.
|
|
207 |
\item @{text "nwq"} is the new waiting queue function. It is calculated using a @{text "case"} statement:
|
|
208 |
\begin{enumerate}
|
|
209 |
\item If the happening event is @{text "P thread cs"}, @{text "thread"} is added to
|
|
210 |
the end of @{text "cs"}'s waiting queue.
|
|
211 |
\item If the happening event is @{text "V thread cs"} and @{text "s"} is a legal state,
|
|
212 |
@{text "th'"} must equal to @{text "thread"},
|
|
213 |
because @{text "thread"} is the one currently holding @{text "cs"}.
|
|
214 |
The case @{text "[] \<Longrightarrow> []"} may never be executed in a legal state.
|
|
215 |
the @{text "(SOME q. distinct q \<and> set q = set qs)"} is used to choose arbitrarily one
|
|
216 |
thread in waiting to take over the released resource @{text "cs"}. In our representation,
|
|
217 |
this amounts to rearrange elements in waiting queue, so that one of them is put at the head.
|
|
218 |
\item For other happening event, the schedule state just does not change.
|
|
219 |
\end{enumerate}
|
|
220 |
\item @{text "ncp"} is new precedence function, it is calculated from the newly updated waiting queue
|
|
221 |
function. The dependency of precedence function on waiting queue function is the reason to
|
|
222 |
put them in the same record so that they can evolve together.
|
|
223 |
\end{enumerate}
|
|
224 |
\end{minipage}
|
|
225 |
*}
|
|
226 |
"schs (e#s) = (let ps = schs s in
|
|
227 |
let pwq = waiting_queue ps in
|
|
228 |
let pcp = cur_preced ps in
|
|
229 |
let nwq = case e of
|
|
230 |
P thread cs \<Rightarrow> pwq(cs:=(pwq cs @ [thread])) |
|
|
231 |
V thread cs \<Rightarrow> let nq = case (pwq cs) of
|
|
232 |
[] \<Rightarrow> [] |
|
|
233 |
(th'#qs) \<Rightarrow> (SOME q. distinct q \<and> set q = set qs)
|
|
234 |
in pwq(cs:=nq) |
|
|
235 |
_ \<Rightarrow> pwq
|
|
236 |
in let ncp = cpreced (e#s) nwq in
|
|
237 |
\<lparr>waiting_queue = nwq, cur_preced = ncp\<rparr>
|
|
238 |
)"
|
|
239 |
|
|
240 |
text {*
|
|
241 |
\noindent
|
|
242 |
The following @{text "wq"} is a shorthand for @{text "waiting_queue"}.
|
|
243 |
*}
|
|
244 |
definition wq :: "state \<Rightarrow> cs \<Rightarrow> thread list"
|
|
245 |
where "wq s = waiting_queue (schs s)"
|
|
246 |
|
|
247 |
text {* \noindent
|
|
248 |
The following @{text "cp"} is a shorthand for @{text "cur_preced"}.
|
|
249 |
*}
|
|
250 |
definition cp :: "state \<Rightarrow> thread \<Rightarrow> precedence"
|
|
251 |
where "cp s = cur_preced (schs s)"
|
|
252 |
|
|
253 |
text {* \noindent
|
|
254 |
Functions @{text "holding"}, @{text "waiting"}, @{text "depend"} and
|
|
255 |
@{text "dependents"} still have the
|
|
256 |
same meaning, but redefined so that they no longer depend on the
|
|
257 |
fictitious {\em waiting queue function}
|
|
258 |
@{text "wq"}, but on system state @{text "s"}.
|
|
259 |
*}
|
|
260 |
defs (overloaded)
|
|
261 |
s_holding_def:
|
|
262 |
"holding (s::state) thread cs \<equiv> (thread \<in> set (wq s cs) \<and> thread = hd (wq s cs))"
|
|
263 |
s_waiting_def:
|
|
264 |
"waiting (s::state) thread cs \<equiv> (thread \<in> set (wq s cs) \<and> thread \<noteq> hd (wq s cs))"
|
|
265 |
s_depend_def:
|
|
266 |
"depend (s::state) \<equiv>
|
|
267 |
{(Th t, Cs c) | t c. waiting (wq s) t c} \<union> {(Cs c, Th t) | c t. holding (wq s) t c}"
|
|
268 |
s_dependents_def:
|
|
269 |
"dependents (s::state) th \<equiv> {th' . (Th th', Th th) \<in> (depend (wq s))^+}"
|
|
270 |
|
|
271 |
text {*
|
|
272 |
The following function @{text "readys"} calculates the set of ready threads. A thread is {\em ready}
|
|
273 |
for running if it is a live thread and it is not waiting for any critical resource.
|
|
274 |
*}
|
|
275 |
definition readys :: "state \<Rightarrow> thread set"
|
|
276 |
where "readys s = {thread . thread \<in> threads s \<and> (\<forall> cs. \<not> waiting s thread cs)}"
|
|
277 |
|
|
278 |
text {* \noindent
|
|
279 |
The following function @{text "runing"} calculates the set of running thread, which is the ready
|
|
280 |
thread with the highest precedence.
|
|
281 |
*}
|
|
282 |
definition runing :: "state \<Rightarrow> thread set"
|
|
283 |
where "runing s = {th . th \<in> readys s \<and> cp s th = Max ((cp s) ` (readys s))}"
|
|
284 |
|
|
285 |
text {* \noindent
|
|
286 |
The following function @{text "holdents s th"} returns the set of resources held by thread
|
|
287 |
@{text "th"} in state @{text "s"}.
|
|
288 |
*}
|
|
289 |
definition holdents :: "state \<Rightarrow> thread \<Rightarrow> cs set"
|
|
290 |
where "holdents s th = {cs . (Cs cs, Th th) \<in> depend s}"
|
|
291 |
|
|
292 |
text {* \noindent
|
|
293 |
@{text "cntCS s th"} returns the number of resources held by thread @{text "th"} in
|
|
294 |
state @{text "s"}:
|
|
295 |
*}
|
|
296 |
definition cntCS :: "state \<Rightarrow> thread \<Rightarrow> nat"
|
|
297 |
where "cntCS s th = card (holdents s th)"
|
|
298 |
|
|
299 |
text {* \noindent
|
|
300 |
The fact that event @{text "e"} is eligible to happen next in state @{text "s"}
|
|
301 |
is expressed as @{text "step s e"}. The predicate @{text "step"} is inductively defined as
|
|
302 |
follows:
|
|
303 |
*}
|
|
304 |
inductive step :: "state \<Rightarrow> event \<Rightarrow> bool"
|
|
305 |
where
|
|
306 |
-- {*
|
|
307 |
A thread can be created if it is not a live thread:
|
|
308 |
*}
|
|
309 |
thread_create: "\<lbrakk>thread \<notin> threads s\<rbrakk> \<Longrightarrow> step s (Create thread prio)" |
|
|
310 |
-- {*
|
|
311 |
A thread can exit if it no longer hold any resource:
|
|
312 |
*}
|
|
313 |
thread_exit: "\<lbrakk>thread \<in> runing s; holdents s thread = {}\<rbrakk> \<Longrightarrow> step s (Exit thread)" |
|
|
314 |
-- {*
|
|
315 |
\begin{minipage}{0.9\textwidth}
|
|
316 |
A thread can request for an critical resource @{text "cs"}, if it is running and
|
|
317 |
the request does not form a loop in the current RAG. The latter condition
|
|
318 |
is set up to avoid deadlock. The condition also reflects our assumption all threads are
|
|
319 |
carefully programmed so that deadlock can not happen:
|
|
320 |
\end{minipage}
|
|
321 |
*}
|
|
322 |
thread_P: "\<lbrakk>thread \<in> runing s; (Cs cs, Th thread) \<notin> (depend s)^+\<rbrakk> \<Longrightarrow>
|
|
323 |
step s (P thread cs)" |
|
|
324 |
-- {*
|
|
325 |
\begin{minipage}{0.9\textwidth}
|
|
326 |
A thread can release a critical resource @{text "cs"}
|
|
327 |
if it is running and holding that resource:
|
|
328 |
\end{minipage}
|
|
329 |
*}
|
|
330 |
thread_V: "\<lbrakk>thread \<in> runing s; holding s thread cs\<rbrakk> \<Longrightarrow> step s (V thread cs)" |
|
|
331 |
-- {*
|
|
332 |
A thread can adjust its own priority as long as it is current running:
|
|
333 |
*}
|
|
334 |
thread_set: "\<lbrakk>thread \<in> runing s\<rbrakk> \<Longrightarrow> step s (Set thread prio)"
|
|
335 |
|
|
336 |
text {* \noindent
|
|
337 |
With predicate @{text "step"}, the fact that @{text "s"} is a legal state in
|
|
338 |
Priority Inheritance protocol can be expressed as: @{text "vt step s"}, where
|
|
339 |
the predicate @{text "vt"} can be defined as the following:
|
|
340 |
*}
|
|
341 |
inductive vt :: "(state \<Rightarrow> event \<Rightarrow> bool) \<Rightarrow> state \<Rightarrow> bool"
|
|
342 |
for cs -- {* @{text "cs"} is an argument representing any step predicate. *}
|
|
343 |
where
|
|
344 |
-- {* Empty list @{text "[]"} is a legal state in any protocol:*}
|
|
345 |
vt_nil[intro]: "vt cs []" |
|
|
346 |
-- {*
|
|
347 |
\begin{minipage}{0.9\textwidth}
|
|
348 |
If @{text "s"} a legal state, and event @{text "e"} is eligible to happen
|
|
349 |
in state @{text "s"}, then @{text "e#s"} is a legal state as well:
|
|
350 |
\end{minipage}
|
|
351 |
*}
|
|
352 |
vt_cons[intro]: "\<lbrakk>vt cs s; cs s e\<rbrakk> \<Longrightarrow> vt cs (e#s)"
|
|
353 |
|
|
354 |
text {* \noindent
|
|
355 |
It is easy to see that the definition of @{text "vt"} is generic. It can be applied to
|
|
356 |
any step predicate to get the set of legal states.
|
|
357 |
*}
|
|
358 |
|
|
359 |
text {* \noindent
|
|
360 |
The following two functions @{text "the_cs"} and @{text "the_th"} are used to extract
|
|
361 |
critical resource and thread respectively out of RAG nodes.
|
|
362 |
*}
|
|
363 |
fun the_cs :: "node \<Rightarrow> cs"
|
|
364 |
where "the_cs (Cs cs) = cs"
|
|
365 |
|
|
366 |
fun the_th :: "node \<Rightarrow> thread"
|
|
367 |
where "the_th (Th th) = th"
|
|
368 |
|
|
369 |
text {* \noindent
|
|
370 |
The following predicate @{text "next_th"} describe the next thread to
|
|
371 |
take over when a critical resource is released. In @{text "next_th s th cs t"},
|
|
372 |
@{text "th"} is the thread to release, @{text "t"} is the one to take over.
|
|
373 |
*}
|
|
374 |
definition next_th:: "state \<Rightarrow> thread \<Rightarrow> cs \<Rightarrow> thread \<Rightarrow> bool"
|
|
375 |
where "next_th s th cs t = (\<exists> rest. wq s cs = th#rest \<and> rest \<noteq> [] \<and>
|
|
376 |
t = hd (SOME q. distinct q \<and> set q = set rest))"
|
|
377 |
|
|
378 |
text {* \noindent
|
|
379 |
The function @{text "count Q l"} is used to count the occurrence of situation @{text "Q"}
|
|
380 |
in list @{text "l"}:
|
|
381 |
*}
|
|
382 |
definition count :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> nat"
|
|
383 |
where "count Q l = length (filter Q l)"
|
|
384 |
|
|
385 |
text {* \noindent
|
|
386 |
The following @{text "cntP s"} returns the number of operation @{text "P"} happened
|
|
387 |
before reaching state @{text "s"}.
|
|
388 |
*}
|
|
389 |
definition cntP :: "state \<Rightarrow> thread \<Rightarrow> nat"
|
|
390 |
where "cntP s th = count (\<lambda> e. \<exists> cs. e = P th cs) s"
|
|
391 |
|
|
392 |
text {* \noindent
|
|
393 |
The following @{text "cntV s"} returns the number of operation @{text "V"} happened
|
|
394 |
before reaching state @{text "s"}.
|
|
395 |
*}
|
|
396 |
definition cntV :: "state \<Rightarrow> thread \<Rightarrow> nat"
|
|
397 |
where "cntV s th = count (\<lambda> e. \<exists> cs. e = V th cs) s"
|
|
398 |
(*<*)
|
|
399 |
end
|
|
400 |
(*>*)
|
|
401 |
|