14499
|
1 |
|
|
2 |
(*<*) theory a2 = Main: (*>*)
|
|
3 |
|
|
4 |
subsection {* Folding Lists and Trees *}
|
|
5 |
|
|
6 |
subsubsection {* Some more list functions *}
|
|
7 |
|
|
8 |
text{* Recall the summation function *}
|
|
9 |
|
|
10 |
(*<*) consts (*>*)
|
|
11 |
sum :: "nat list \<Rightarrow> nat"
|
|
12 |
primrec
|
|
13 |
"sum [] = 0"
|
|
14 |
"sum (x # xs) = x + sum xs"
|
|
15 |
|
|
16 |
|
|
17 |
text{* In the Isabelle library, you will find in theory
|
|
18 |
\texttt{List.thy} the functions @{text foldr} and @{text foldl}, which
|
|
19 |
allow you to define some list functions, among them @{text sum} and
|
|
20 |
@{text length}. Show the following: *}
|
|
21 |
|
|
22 |
|
|
23 |
lemma sum_foldr: "sum xs = foldr (op +) xs 0"
|
|
24 |
(*<*) oops (*>*)
|
|
25 |
|
|
26 |
lemma length_foldr: "length xs = foldr (\<lambda> x res. 1 + res) xs 0"
|
|
27 |
(*<*) oops (*>*)
|
|
28 |
|
|
29 |
text {* Repeated application of @{text foldr} and @{text map} has the
|
|
30 |
disadvantage that a list is traversed several times. A single traversal is sufficient, as
|
|
31 |
illustrated by the following example: *}
|
|
32 |
|
|
33 |
lemma "sum (map (\<lambda> x. x + 3) xs) = foldr h xs b"
|
|
34 |
(*<*) oops (*>*)
|
|
35 |
|
|
36 |
text {* Find terms @{text h} and @{text b} which solve this
|
|
37 |
equation. Generalize this result, i.e. show for appropriate @{text h}
|
|
38 |
and @{text b}: *}
|
|
39 |
|
|
40 |
|
|
41 |
lemma "foldr g (map f xs) a = foldr h xs b"
|
|
42 |
(*<*) oops (*>*)
|
|
43 |
|
|
44 |
text {* Hint: Isabelle can help you find the solution if you use the
|
|
45 |
equalities arising during a proof attempt. *}
|
|
46 |
|
|
47 |
text {* The following function @{text rev_acc} reverses a list in linear time: *}
|
|
48 |
|
|
49 |
|
|
50 |
consts
|
|
51 |
rev_acc :: "['a list, 'a list] \<Rightarrow> 'a list"
|
|
52 |
primrec
|
|
53 |
"rev_acc [] ys = ys"
|
|
54 |
"rev_acc (x#xs) ys = (rev_acc xs (x#ys))"
|
|
55 |
|
|
56 |
text{* Show that @{text rev_acc} can be defined by means of @{text foldl}. *}
|
|
57 |
|
|
58 |
lemma rev_acc_foldl: "rev_acc xs a = foldl (\<lambda> ys x. x # ys) a xs"
|
|
59 |
(*<*) oops (*>*)
|
|
60 |
|
|
61 |
text {* On the first exercise sheet, we have shown: *}
|
|
62 |
|
|
63 |
lemma sum_append [simp]: "sum (xs @ ys) = sum xs + sum ys"
|
|
64 |
by (induct xs) auto
|
|
65 |
|
|
66 |
text {* Prove a similar distributivity property for @{text foldr},
|
|
67 |
i.e. something like @{text "foldr f (xs @ ys) a = f (foldr f xs a)
|
|
68 |
(foldr f ys a)"}. However, you will have to strengthen the premisses
|
|
69 |
by taking into account algebraic properties of @{text f} and @{text
|
|
70 |
a}. *}
|
|
71 |
|
|
72 |
|
|
73 |
lemma foldr_append: "foldr f (xs @ ys) a = f (foldr f xs a) (foldr f ys a)"
|
|
74 |
(*<*) oops (*>*)
|
|
75 |
|
|
76 |
text {* Now, define the function @{text prod}, which computes the product of all list elements: *}
|
|
77 |
(*<*) consts (*>*)
|
|
78 |
prod :: "nat list \<Rightarrow> nat"
|
|
79 |
|
|
80 |
text {* direcly with the aid of a fold and prove the following: *}
|
|
81 |
lemma "prod (xs @ ys) = prod xs * prod ys"
|
|
82 |
(*<*) oops (*>*)
|
|
83 |
|
|
84 |
|
|
85 |
subsubsection {* Functions on Trees *}
|
|
86 |
|
|
87 |
text {* Consider the following type of binary trees: *}
|
|
88 |
datatype 'a tree = Tip | Node "'a tree" 'a "'a tree"
|
|
89 |
|
|
90 |
text {* Define functions which convert a tree into a list by traversing it in pre- resp. postorder: *}
|
|
91 |
(*<*) consts (*>*)
|
|
92 |
preorder :: "'a tree \<Rightarrow> 'a list"
|
|
93 |
postorder :: "'a tree \<Rightarrow> 'a list"
|
|
94 |
|
|
95 |
text {* You have certainly realized that computation of postorder traversal can be efficiently realized with an accumulator, in analogy to @{text rev_acc}: *}
|
|
96 |
|
|
97 |
consts
|
|
98 |
postorder_acc :: "['a tree, 'a list] \<Rightarrow> 'a list"
|
|
99 |
|
|
100 |
text {* Define this function and show: *}
|
|
101 |
lemma "postorder_acc t xs = (postorder t) @ xs"
|
|
102 |
(*<*) oops (*>*)
|
|
103 |
|
|
104 |
|
|
105 |
text {* @{text postorder_acc} is the instance of a function
|
|
106 |
@{text foldl_tree}, which is similar to @{text foldl}. *}
|
|
107 |
|
|
108 |
consts
|
|
109 |
foldl_tree :: "('b => 'a => 'b) \<Rightarrow> 'b \<Rightarrow> 'a tree \<Rightarrow> 'b"
|
|
110 |
|
|
111 |
text {* Show the following: *}
|
|
112 |
|
|
113 |
lemma "\<forall> a. postorder_acc t a = foldl_tree (\<lambda> xs x. Cons x xs) a t"
|
|
114 |
(*<*) oops (*>*)
|
|
115 |
|
|
116 |
text {* Define a function @{text tree_sum} that computes the sum of
|
|
117 |
the elements of a tree of natural numbers: *}
|
|
118 |
|
|
119 |
consts
|
|
120 |
tree_sum :: "nat tree \<Rightarrow> nat"
|
|
121 |
|
|
122 |
text {* and show that this function satisfies *}
|
|
123 |
|
|
124 |
lemma "tree_sum t = sum (preorder t)"
|
|
125 |
(*<*) oops (*>*)
|
|
126 |
|
|
127 |
|
|
128 |
(*<*) end (*>*)
|
|
129 |
|