author | nipkow |
Sat, 05 Feb 2005 19:24:11 +0100 | |
changeset 15500 | dd4ab096f082 |
parent 14981 | e73f8140af78 |
child 16078 | e1364521a250 |
permissions | -rw-r--r-- |
7018 | 1 |
(* Title: HOL/Induct/Tree.thy |
2 |
ID: $Id$ |
|
3 |
Author: Stefan Berghofer, TU Muenchen |
|
4 |
*) |
|
5 |
||
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
7018
diff
changeset
|
6 |
header {* Infinitely branching trees *} |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
7018
diff
changeset
|
7 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
7018
diff
changeset
|
8 |
theory Tree = Main: |
7018 | 9 |
|
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
7018
diff
changeset
|
10 |
datatype 'a tree = |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
7018
diff
changeset
|
11 |
Atom 'a |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
7018
diff
changeset
|
12 |
| Branch "nat => 'a tree" |
7018 | 13 |
|
14 |
consts |
|
15 |
map_tree :: "('a => 'b) => 'a tree => 'b tree" |
|
16 |
primrec |
|
17 |
"map_tree f (Atom a) = Atom (f a)" |
|
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
7018
diff
changeset
|
18 |
"map_tree f (Branch ts) = Branch (\<lambda>x. map_tree f (ts x))" |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
7018
diff
changeset
|
19 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
7018
diff
changeset
|
20 |
lemma tree_map_compose: "map_tree g (map_tree f t) = map_tree (g \<circ> f) t" |
12171 | 21 |
by (induct t) simp_all |
7018 | 22 |
|
23 |
consts |
|
24 |
exists_tree :: "('a => bool) => 'a tree => bool" |
|
25 |
primrec |
|
26 |
"exists_tree P (Atom a) = P a" |
|
11046
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
7018
diff
changeset
|
27 |
"exists_tree P (Branch ts) = (\<exists>x. exists_tree P (ts x))" |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
7018
diff
changeset
|
28 |
|
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
7018
diff
changeset
|
29 |
lemma exists_map: |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
7018
diff
changeset
|
30 |
"(!!x. P x ==> Q (f x)) ==> |
b5f5942781a0
Induct: converted some theories to new-style format;
wenzelm
parents:
7018
diff
changeset
|
31 |
exists_tree P ts ==> exists_tree Q (map_tree f ts)" |
12171 | 32 |
by (induct ts) auto |
7018 | 33 |
|
34 |
end |