Diferencia entre revisiones de «Relación 11»
De Razonamiento automático (2014-15)
(Página creada con '<source lang="isar"> header {* R11: Plegados de listas y de árboles *} theory R11 imports Main begin section {* Nuevas funciones sobre listas *} text {* Nota. En esta r...') |
|||
Línea 1: | Línea 1: | ||
<source lang="isar"> | <source lang="isar"> | ||
header {* R11: Plegados de listas y de árboles *} | header {* R11: Plegados de listas y de árboles *} | ||
− | + | ||
theory R11 | theory R11 | ||
imports Main | imports Main | ||
begin | begin | ||
− | + | ||
section {* Nuevas funciones sobre listas *} | section {* Nuevas funciones sobre listas *} | ||
− | + | ||
text {* | text {* | ||
Nota. En esta relación se usará la función suma tal que (suma xs) es | Nota. En esta relación se usará la función suma tal que (suma xs) es | ||
Línea 13: | Línea 13: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
+ | fun suma1 :: "nat list ⇒ nat" where | ||
+ | "suma1 xs = foldr (op +) xs 0" | ||
+ | |||
+ | --"jeshorcob" | ||
fun suma :: "nat list ⇒ nat" where | fun suma :: "nat list ⇒ nat" where | ||
− | "suma xs = | + | "suma [] = 0" |
− | + | |"suma (x#xs) = x+suma xs" | |
+ | |||
text {* | text {* | ||
Las funciones de plegado, foldr y foldl, están definidas en la teoría | Las funciones de plegado, foldr y foldl, están definidas en la teoría | ||
Línea 22: | Línea 28: | ||
foldr f [] = id | foldr f [] = id | ||
foldr f (x # xs) = f x ∘ foldr f xs | foldr f (x # xs) = f x ∘ foldr f xs | ||
− | + | ||
foldl :: "('b ⇒ 'a ⇒ 'b) ⇒ 'b ⇒ 'a list ⇒ 'b" | foldl :: "('b ⇒ 'a ⇒ 'b) ⇒ 'b ⇒ 'a list ⇒ 'b" | ||
foldl f a [] = a | foldl f a [] = a | ||
foldl f a (x # xs) = foldl f (f a x) xs" | foldl f a (x # xs) = foldl f (f a x) xs" | ||
− | + | ||
Por ejemplo, | Por ejemplo, | ||
foldr (op +) [a,b,c] d = a + (b + (c + d)) | foldr (op +) [a,b,c] d = a + (b + (c + d)) | ||
Línea 33: | Línea 39: | ||
foldl (op -) (0::int) [9,4,2] = -15 | foldl (op -) (0::int) [9,4,2] = -15 | ||
*} | *} | ||
− | + | ||
value "foldr (op +) [a,b,c] d" -- "= a + (b + (c + d))" | value "foldr (op +) [a,b,c] d" -- "= a + (b + (c + d))" | ||
value "foldl (op +) d [a,b,c]" -- "= ((d + a) + b) + c" | value "foldl (op +) d [a,b,c]" -- "= ((d + a) + b) + c" | ||
value "foldr (op -) [9,4,2] (0::int)" -- "= 7" | value "foldr (op -) [9,4,2] (0::int)" -- "= 7" | ||
value "foldl (op -) (0::int) [9,4,2]" -- "= -15" | value "foldl (op -) (0::int) [9,4,2]" -- "= -15" | ||
− | + | ||
text {* | text {* | ||
--------------------------------------------------------------------- | --------------------------------------------------------------------- | ||
Línea 46: | Línea 52: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
lemma suma_foldr: "suma xs = foldr (op +) xs 0" | lemma suma_foldr: "suma xs = foldr (op +) xs 0" | ||
− | + | by (induct xs, simp_all) | |
text {* | text {* | ||
Línea 56: | Línea 63: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
lemma length_foldr: "length xs = foldr (λ x res. 1 + res) xs 0" | lemma length_foldr: "length xs = foldr (λ x res. 1 + res) xs 0" | ||
− | + | by (induct xs, simp_all) | |
− | + | ||
text {* | text {* | ||
--------------------------------------------------------------------- | --------------------------------------------------------------------- | ||
Línea 70: | Línea 78: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
+ | lemma "suma (map (λx. x + 3) xs) = foldr (λ x y. x+y+3) xs 0" | ||
+ | by (induct xs, simp_all) | ||
text {* | text {* | ||
Línea 79: | Línea 90: | ||
--------------------------------------------------------------------- | --------------------------------------------------------------------- | ||
*} | *} | ||
+ | |||
+ | --"jeshorcob" | ||
+ | lemma "foldr g (map f xs) a = foldr (g∘f) xs a" | ||
+ | by (induct xs, simp_all) | ||
+ | |||
text {* | text {* | ||
Línea 86: | Línea 102: | ||
"inversa_ac [] ys = ys" | "inversa_ac [] ys = ys" | ||
| "inversa_ac (x#xs) ys = (inversa_ac xs (x#ys))" | | "inversa_ac (x#xs) ys = (inversa_ac xs (x#ys))" | ||
− | + | ||
definition inversa_ac :: "'a list ⇒ 'a list" where | definition inversa_ac :: "'a list ⇒ 'a list" where | ||
"inversa_ac xs ≡ inversa_ac_aux xs []" | "inversa_ac xs ≡ inversa_ac_aux xs []" | ||
Línea 95: | Línea 111: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
fun inversa_ac_aux :: "['a list, 'a list] ⇒ 'a list" where | fun inversa_ac_aux :: "['a list, 'a list] ⇒ 'a list" where | ||
"inversa_ac_aux [] ys = ys" | "inversa_ac_aux [] ys = ys" | ||
| "inversa_ac_aux (x#xs) ys = (inversa_ac_aux xs (x#ys))" | | "inversa_ac_aux (x#xs) ys = (inversa_ac_aux xs (x#ys))" | ||
+ | --"jeshorcob" | ||
definition inversa_ac :: "'a list ⇒ 'a list" where | definition inversa_ac :: "'a list ⇒ 'a list" where | ||
"inversa_ac xs ≡ inversa_ac_aux xs []" | "inversa_ac xs ≡ inversa_ac_aux xs []" | ||
− | + | ||
value "inversa_ac [a,d,b,c]" -- "= [c, b, d, a]" | value "inversa_ac [a,d,b,c]" -- "= [c, b, d, a]" | ||
+ | --"jeshorcob" | ||
lemma inversa_ac_aux_foldl: | lemma inversa_ac_aux_foldl: | ||
− | "inversa_ac_aux xs a = foldl | + | "inversa_ac_aux xs a = foldl (λ ys x. x#ys) a xs" |
− | + | by (induct xs arbitrary: a, simp_all) | |
text {* | text {* | ||
Línea 116: | Línea 135: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
lemma suma_append [simp]: | lemma suma_append [simp]: | ||
"suma (xs @ ys) = suma xs + suma ys" | "suma (xs @ ys) = suma xs + suma ys" | ||
− | + | by (induct xs, simp_all) | |
text {* | text {* | ||
Línea 128: | Línea 148: | ||
--------------------------------------------------------------------- | --------------------------------------------------------------------- | ||
*} | *} | ||
+ | |||
+ | --"jeshorcob" | ||
+ | lemma l1: | ||
+ | assumes a1: "(∀ x. (f a x = x))" and | ||
+ | a2: "(∀x y z. f (f x y) z = f x (f y z))" | ||
+ | shows "foldr f (xs @ ys) a = f (foldr f xs a) (foldr f ys a)" | ||
+ | proof (induct xs) | ||
+ | show "foldr f ([] @ ys) a = f (foldr f [] a) (foldr f ys a)" | ||
+ | using a1 by simp | ||
+ | next | ||
+ | show "⋀aa xs. | ||
+ | foldr f (xs @ ys) a = f (foldr f xs a) (foldr f ys a) ⟹ | ||
+ | foldr f ((aa # xs) @ ys) a = f (foldr f (aa # xs) a) (foldr f ys a)" | ||
+ | using a2 by simp | ||
+ | qed | ||
+ | |||
+ | (*jeshorcob: | ||
+ | saco fuera las hipotesis porque se usan tambien en el ejercicio siguiente, | ||
+ | ademas la prueba queda mas corta *) | ||
+ | |||
+ | --"jeshorcob" | ||
+ | definition a1 :: "['a ⇒ 'b ⇒ 'b, 'a] ⇒ bool" where | ||
+ | "a1 f a == (∀ x. (f a x = x))" | ||
+ | |||
+ | --"jeshorcob" | ||
+ | definition a2 :: "['a ⇒ 'a ⇒ 'a] ⇒ bool" where | ||
+ | "a2 f == (∀ x y z. f (f x y) z = f x (f y z))" | ||
+ | |||
+ | --"jeshorcob" | ||
+ | lemma l2: "⟦a1 f a ; a2 f⟧ | ||
+ | ⟹ foldr f (xs @ ys) a = f (foldr f xs a) (foldr f ys a)" | ||
+ | by (induct xs, simp add: a1_def, simp add: a2_def) | ||
+ | |||
text {* | text {* | ||
Línea 137: | Línea 190: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
definition prod :: "nat list ⇒ nat" where | definition prod :: "nat list ⇒ nat" where | ||
− | "prod xs ≡ | + | "prod xs ≡ foldr (op *) xs 1" |
− | + | ||
− | value "prod [2::nat,3,5]" -- " | + | value "prod [2::nat,3,5] = 30" -- "True" |
− | + | ||
text {* | text {* | ||
--------------------------------------------------------------------- | --------------------------------------------------------------------- | ||
Línea 149: | Línea 203: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
lemma "prod (xs @ ys) = prod xs * prod ys" | lemma "prod (xs @ ys) = prod xs * prod ys" | ||
− | + | by (simp only: prod_def, rule l2, simp add: a1_def, simp add: a2_def) | |
section {* Functiones sobre árboles *} | section {* Functiones sobre árboles *} | ||
− | + | ||
text {* | text {* | ||
--------------------------------------------------------------------- | --------------------------------------------------------------------- | ||
Línea 168: | Línea 223: | ||
--------------------------------------------------------------------- | --------------------------------------------------------------------- | ||
*} | *} | ||
− | + | ||
datatype 'a arbol = H | N "'a arbol" "'a" "'a arbol" | datatype 'a arbol = H | N "'a arbol" "'a" "'a arbol" | ||
− | + | ||
value "N (N H c H) e (N H g H)" | value "N (N H c H) e (N H g H)" | ||
− | + | ||
text {* | text {* | ||
--------------------------------------------------------------------- | --------------------------------------------------------------------- | ||
Línea 184: | Línea 239: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
fun preOrden :: "'a arbol ⇒ 'a list" where | fun preOrden :: "'a arbol ⇒ 'a list" where | ||
− | "preOrden | + | "preOrden H = []" |
− | + | |"preOrden (N i x d) = x#(preOrden i @ preOrden d)" | |
− | value "preOrden (N (N H c H) e (N H g H)) | + | |
− | + | value "preOrden (N (N H c H) e (N H g H)) = [e,c,g]" -- "True" | |
− | + | ||
text {* | text {* | ||
--------------------------------------------------------------------- | --------------------------------------------------------------------- | ||
Línea 201: | Línea 257: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
fun postOrden :: "'a arbol ⇒ 'a list" where | fun postOrden :: "'a arbol ⇒ 'a list" where | ||
− | "postOrden | + | "postOrden H = []" |
− | + | |"postOrden (N i x d) = (postOrden i) @ (postOrden d) @ [x]" | |
− | value "postOrden (N (N H c H) e (N H g H)) | + | |
− | + | value "postOrden (N (N H c H) e (N H g H)) = [c,g,e]" -- "True" | |
− | + | ||
text {* | text {* | ||
--------------------------------------------------------------------- | --------------------------------------------------------------------- | ||
Línea 218: | Línea 275: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
fun postOrdenAaux :: "['a arbol, 'a list] ⇒ 'a list" where | fun postOrdenAaux :: "['a arbol, 'a list] ⇒ 'a list" where | ||
− | "postOrdenAaux | + | "postOrdenAaux H xs = xs" |
+ | |"postOrdenAaux (N i x d) xs = (postOrdenAaux i (postOrdenAaux d (x#xs)))" | ||
+ | --"jeshorcob" | ||
definition postOrdenA :: "'a arbol ⇒ 'a list" where | definition postOrdenA :: "'a arbol ⇒ 'a list" where | ||
− | "postOrdenA a ≡ | + | "postOrdenA a ≡ postOrdenAaux a []" |
− | + | ||
− | value "postOrdenA (N (N H c H) e (N H g H)) | + | value "postOrdenA (N (N H c H) e (N H g H)) = [c,g,e]" -- "True" |
− | + | ||
− | |||
text {* | text {* | ||
--------------------------------------------------------------------- | --------------------------------------------------------------------- | ||
Línea 234: | Línea 293: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
lemma "postOrdenAaux a xs = (postOrden a) @ xs" | lemma "postOrdenAaux a xs = (postOrden a) @ xs" | ||
− | + | by (induct a arbitrary: xs, simp_all) | |
− | + | ||
text {* | text {* | ||
--------------------------------------------------------------------- | --------------------------------------------------------------------- | ||
Línea 246: | Línea 306: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
fun foldl_arbol :: "('b => 'a => 'b) ⇒ 'b ⇒ 'a arbol ⇒ 'b" where | fun foldl_arbol :: "('b => 'a => 'b) ⇒ 'b ⇒ 'a arbol ⇒ 'b" where | ||
− | "foldl_arbol f b | + | "foldl_arbol f b H = b" |
+ | |"foldl_arbol f b (N i x d) = foldl_arbol f (foldl_arbol f (f b x) d) i" | ||
text {* | text {* | ||
Línea 256: | Línea 318: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
lemma "postOrdenAaux t a = foldl_arbol (λ xs x. Cons x xs) a t" | lemma "postOrdenAaux t a = foldl_arbol (λ xs x. Cons x xs) a t" | ||
− | + | by (induct t arbitrary: a, simp_all) | |
− | + | ||
text {* | text {* | ||
--------------------------------------------------------------------- | --------------------------------------------------------------------- | ||
Línea 268: | Línea 331: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
fun suma_arbol :: "nat arbol ⇒ nat" where | fun suma_arbol :: "nat arbol ⇒ nat" where | ||
− | "suma_arbol | + | "suma_arbol H = 0" |
− | + | |"suma_arbol (N i x d) = (suma_arbol i) + x + (suma_arbol d)" | |
+ | |||
text {* | text {* | ||
--------------------------------------------------------------------- | --------------------------------------------------------------------- | ||
Línea 278: | Línea 343: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
lemma "suma_arbol a = suma (preOrden a)" | lemma "suma_arbol a = suma (preOrden a)" | ||
− | + | by (induct a, simp_all) | |
− | + | ||
end | end | ||
</source> | </source> |
Revisión del 13:12 14 ene 2015
header {* R11: Plegados de listas y de árboles *}
theory R11
imports Main
begin
section {* Nuevas funciones sobre listas *}
text {*
Nota. En esta relación se usará la función suma tal que (suma xs) es
la suma de los elementos de xs, definida por
*}
--"jeshorcob"
fun suma1 :: "nat list ⇒ nat" where
"suma1 xs = foldr (op +) xs 0"
--"jeshorcob"
fun suma :: "nat list ⇒ nat" where
"suma [] = 0"
|"suma (x#xs) = x+suma xs"
text {*
Las funciones de plegado, foldr y foldl, están definidas en la teoría
List.thy por
foldr :: "('a ⇒ 'b ⇒ 'b) ⇒ 'a list ⇒ 'b ⇒ 'b"
foldr f [] = id
foldr f (x # xs) = f x ∘ foldr f xs
foldl :: "('b ⇒ 'a ⇒ 'b) ⇒ 'b ⇒ 'a list ⇒ 'b"
foldl f a [] = a
foldl f a (x # xs) = foldl f (f a x) xs"
Por ejemplo,
foldr (op +) [a,b,c] d = a + (b + (c + d))
foldl (op +) d [a,b,c] = ((d + a) + b) + c
foldr (op -) [9,4,2] (0::int) = 7
foldl (op -) (0::int) [9,4,2] = -15
*}
value "foldr (op +) [a,b,c] d" -- "= a + (b + (c + d))"
value "foldl (op +) d [a,b,c]" -- "= ((d + a) + b) + c"
value "foldr (op -) [9,4,2] (0::int)" -- "= 7"
value "foldl (op -) (0::int) [9,4,2]" -- "= -15"
text {*
---------------------------------------------------------------------
Ejercicio 1. Demostrar que
suma xs = foldr (op +) xs 0
---------------------------------------------------------------------
*}
--"jeshorcob"
lemma suma_foldr: "suma xs = foldr (op +) xs 0"
by (induct xs, simp_all)
text {*
---------------------------------------------------------------------
Ejercicio 2. Demostrar que
length xs = foldr (λ x res. 1 + res) xs 0
---------------------------------------------------------------------
*}
--"jeshorcob"
lemma length_foldr: "length xs = foldr (λ x res. 1 + res) xs 0"
by (induct xs, simp_all)
text {*
---------------------------------------------------------------------
Ejercicio 3. La aplicación repetida de foldr y map tiene el
inconveniente de que la lista se recorre varias veces. Sin embargo, es
suficiente recorrerla una vez como se muestra en el siguiente ejemplo,
suma (map (λx. x + 3) xs) = foldr h xs b
Determinar los valores de h y b para que se verifique la igualdad
anterior y demostrarla.
---------------------------------------------------------------------
*}
--"jeshorcob"
lemma "suma (map (λx. x + 3) xs) = foldr (λ x y. x+y+3) xs 0"
by (induct xs, simp_all)
text {*
---------------------------------------------------------------------
Ejercicio 4. Generalizar el resultado anterior; es decir determinar
los valores de h y b para que se verifique la igualdad
foldr g (map f xs) a = foldr h xs b
y demostrarla.
---------------------------------------------------------------------
*}
--"jeshorcob"
lemma "foldr g (map f xs) a = foldr (g∘f) xs a"
by (induct xs, simp_all)
text {*
---------------------------------------------------------------------
Ejercicio 5. La siguiente función invierte una lista en tiempo lineal
fun inversa_ac :: "['a list, 'a list] ⇒ 'a list" where
"inversa_ac [] ys = ys"
| "inversa_ac (x#xs) ys = (inversa_ac xs (x#ys))"
definition inversa_ac :: "'a list ⇒ 'a list" where
"inversa_ac xs ≡ inversa_ac_aux xs []"
Por ejemplo,
inversa_ac [a,d,b,c] = [c, b, d, a]
Demostrar que inversa_ac se puede definir unsando foldl.
---------------------------------------------------------------------
*}
--"jeshorcob"
fun inversa_ac_aux :: "['a list, 'a list] ⇒ 'a list" where
"inversa_ac_aux [] ys = ys"
| "inversa_ac_aux (x#xs) ys = (inversa_ac_aux xs (x#ys))"
--"jeshorcob"
definition inversa_ac :: "'a list ⇒ 'a list" where
"inversa_ac xs ≡ inversa_ac_aux xs []"
value "inversa_ac [a,d,b,c]" -- "= [c, b, d, a]"
--"jeshorcob"
lemma inversa_ac_aux_foldl:
"inversa_ac_aux xs a = foldl (λ ys x. x#ys) a xs"
by (induct xs arbitrary: a, simp_all)
text {*
---------------------------------------------------------------------
Ejercicio 6. Demostrar la siguiente propiedad distributiva de la suma
sobre la concatenación:
suma (xs @ ys) = suma xs + suma ys
---------------------------------------------------------------------
*}
--"jeshorcob"
lemma suma_append [simp]:
"suma (xs @ ys) = suma xs + suma ys"
by (induct xs, simp_all)
text {*
---------------------------------------------------------------------
Ejercicio 7. Demostrar una propiedad similar para foldr
foldr f (xs @ ys) a = f (foldr f xs a) (foldr f ys a)
En este caso, hay que restringir el resultado teniendo en cuenta
propiedades algebraicas de f y a.
---------------------------------------------------------------------
*}
--"jeshorcob"
lemma l1:
assumes a1: "(∀ x. (f a x = x))" and
a2: "(∀x y z. f (f x y) z = f x (f y z))"
shows "foldr f (xs @ ys) a = f (foldr f xs a) (foldr f ys a)"
proof (induct xs)
show "foldr f ([] @ ys) a = f (foldr f [] a) (foldr f ys a)"
using a1 by simp
next
show "⋀aa xs.
foldr f (xs @ ys) a = f (foldr f xs a) (foldr f ys a) ⟹
foldr f ((aa # xs) @ ys) a = f (foldr f (aa # xs) a) (foldr f ys a)"
using a2 by simp
qed
(*jeshorcob:
saco fuera las hipotesis porque se usan tambien en el ejercicio siguiente,
ademas la prueba queda mas corta *)
--"jeshorcob"
definition a1 :: "['a ⇒ 'b ⇒ 'b, 'a] ⇒ bool" where
"a1 f a == (∀ x. (f a x = x))"
--"jeshorcob"
definition a2 :: "['a ⇒ 'a ⇒ 'a] ⇒ bool" where
"a2 f == (∀ x y z. f (f x y) z = f x (f y z))"
--"jeshorcob"
lemma l2: "⟦a1 f a ; a2 f⟧
⟹ foldr f (xs @ ys) a = f (foldr f xs a) (foldr f ys a)"
by (induct xs, simp add: a1_def, simp add: a2_def)
text {*
---------------------------------------------------------------------
Ejercicio 8. Definir, usando foldr, la función
prod :: "nat list ⇒ nat"
tal que (prod xs) es el producto de los elementos de xs. Por ejemplo,
---------------------------------------------------------------------
*}
--"jeshorcob"
definition prod :: "nat list ⇒ nat" where
"prod xs ≡ foldr (op *) xs 1"
value "prod [2::nat,3,5] = 30" -- "True"
text {*
---------------------------------------------------------------------
Ejercicio 9. Demostrar directamente (es decir, sin inducción) que
prod (xs @ ys) = prod xs * prod ys
---------------------------------------------------------------------
*}
--"jeshorcob"
lemma "prod (xs @ ys) = prod xs * prod ys"
by (simp only: prod_def, rule l2, simp add: a1_def, simp add: a2_def)
section {* Functiones sobre árboles *}
text {*
---------------------------------------------------------------------
Ejercicio 10. Definir el tipo de datos arbol para representar los
árboles binarios que tiene información sólo en los nodos. Por ejemplo,
el árbol
e
/ \
/ \
c g
/ \ / \
· · · ·
se representa por "N (N H c H) e (N H g H)"
---------------------------------------------------------------------
*}
datatype 'a arbol = H | N "'a arbol" "'a" "'a arbol"
value "N (N H c H) e (N H g H)"
text {*
---------------------------------------------------------------------
Ejercicio 11. Definir la función
preOrden :: "'a arbol ⇒ 'a list"
tal que (preOrden a) es el recorrido pre orden del árbol a. Por
ejemplo,
preOrden (N (N H c H) e (N H g H))
= [e,c,g]
---------------------------------------------------------------------
*}
--"jeshorcob"
fun preOrden :: "'a arbol ⇒ 'a list" where
"preOrden H = []"
|"preOrden (N i x d) = x#(preOrden i @ preOrden d)"
value "preOrden (N (N H c H) e (N H g H)) = [e,c,g]" -- "True"
text {*
---------------------------------------------------------------------
Ejercicio 12. Definir la función
postOrden :: "'a arbol ⇒ 'a list"
tal que (postOrden a) es el recorrido post orden del árbol a. Por
ejemplo,
postOrden (N (N H c H) e (N H g H))
= [c,g,e]
---------------------------------------------------------------------
*}
--"jeshorcob"
fun postOrden :: "'a arbol ⇒ 'a list" where
"postOrden H = []"
|"postOrden (N i x d) = (postOrden i) @ (postOrden d) @ [x]"
value "postOrden (N (N H c H) e (N H g H)) = [c,g,e]" -- "True"
text {*
---------------------------------------------------------------------
Ejercicio 13. Definir, usando un acumulador, la función
postOrdenA :: "'a arbol ⇒ 'a list"
tal que (postOrdenA a) es el recorrido post orden del árbol a. Por
ejemplo,
postOrdenA (N (N H c H) e (N H g H))
= [c,g,e]
---------------------------------------------------------------------
*}
--"jeshorcob"
fun postOrdenAaux :: "['a arbol, 'a list] ⇒ 'a list" where
"postOrdenAaux H xs = xs"
|"postOrdenAaux (N i x d) xs = (postOrdenAaux i (postOrdenAaux d (x#xs)))"
--"jeshorcob"
definition postOrdenA :: "'a arbol ⇒ 'a list" where
"postOrdenA a ≡ postOrdenAaux a []"
value "postOrdenA (N (N H c H) e (N H g H)) = [c,g,e]" -- "True"
text {*
---------------------------------------------------------------------
Ejercicio 14. Demostrar que
postOrdenAaux a xs = (postOrden a) @ xs
---------------------------------------------------------------------
*}
--"jeshorcob"
lemma "postOrdenAaux a xs = (postOrden a) @ xs"
by (induct a arbitrary: xs, simp_all)
text {*
---------------------------------------------------------------------
Ejercicio 15. Definir la función
foldl_arbol :: "('b => 'a => 'b) ⇒ 'b ⇒ 'a arbol ⇒ 'b" where
tal que (foldl_arbol f b a) es el plegado izquierdo del árbol a con la
operación f y elemento inicial b.
---------------------------------------------------------------------
*}
--"jeshorcob"
fun foldl_arbol :: "('b => 'a => 'b) ⇒ 'b ⇒ 'a arbol ⇒ 'b" where
"foldl_arbol f b H = b"
|"foldl_arbol f b (N i x d) = foldl_arbol f (foldl_arbol f (f b x) d) i"
text {*
---------------------------------------------------------------------
Ejercicio 16. Demostrar que
postOrdenAaux t a = foldl_arbol (λ xs x. Cons x xs) a t
---------------------------------------------------------------------
*}
--"jeshorcob"
lemma "postOrdenAaux t a = foldl_arbol (λ xs x. Cons x xs) a t"
by (induct t arbitrary: a, simp_all)
text {*
---------------------------------------------------------------------
Ejercicio 17. Definir la función
suma_arbol :: "nat arbol ⇒ nat"
tal que (suma_arbol a) es la suma de los elementos del árbol de
números naturales a.
---------------------------------------------------------------------
*}
--"jeshorcob"
fun suma_arbol :: "nat arbol ⇒ nat" where
"suma_arbol H = 0"
|"suma_arbol (N i x d) = (suma_arbol i) + x + (suma_arbol d)"
text {*
---------------------------------------------------------------------
Ejercicio 18. Demostrar que
suma_arbol a = suma (preOrden a)"
---------------------------------------------------------------------
*}
--"jeshorcob"
lemma "suma_arbol a = suma (preOrden a)"
by (induct a, simp_all)
end