author | huffman |
Fri, 12 Aug 2011 14:45:50 -0700 | |
changeset 44174 | d1d79f0e1ea6 |
parent 43158 | 686fa0a0696e |
child 44177 | b4b5cbca2519 |
permissions | -rw-r--r-- |
43158 | 1 |
(* Author: Tobias Nipkow *) |
2 |
||
3 |
theory Util imports Main |
|
4 |
begin |
|
5 |
||
6 |
subsection "Show sets of variables as lists" |
|
7 |
||
8 |
text {* Sets are functions and are not displayed by element if |
|
9 |
computed as values: *} |
|
10 |
value "{''x'', ''y''}" |
|
11 |
||
12 |
text {* Lists do not have this problem *} |
|
13 |
value "[''x'', ''y'']" |
|
14 |
||
15 |
text {* We define a function @{text show} that converts a set of |
|
16 |
variable names into a list. To keep things simple, we rely on |
|
17 |
the convention that we only use single letter names. |
|
18 |
*} |
|
19 |
definition |
|
20 |
letters :: "string list" where |
|
21 |
"letters = map (\<lambda>c. [c]) chars" |
|
22 |
||
23 |
definition |
|
24 |
"show" :: "string set \<Rightarrow> string list" where |
|
44174
d1d79f0e1ea6
make more HOL theories work with separate set type
huffman
parents:
43158
diff
changeset
|
25 |
"show S = filter (\<lambda>x. x \<in> S) letters" |
43158 | 26 |
|
27 |
value "show {''x'', ''z''}" |
|
28 |
||
29 |
end |