Diferencia entre revisiones de «Relación 7»
De Razonamiento automático (2014-15)
(Página creada con '<source lang="isar"> header {* R7: Recorridos de árboles *} theory R7 imports Main begin text {* --------------------------------------------------------------------- ...') |
|||
Línea 36: | Línea 36: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
fun preOrden :: "'a arbol ⇒ 'a list" where | fun preOrden :: "'a arbol ⇒ 'a list" where | ||
− | "preOrden | + | "preOrden (H x) = [x]" |
+ | |"preOrden (N x yy zz) = x#(preOrden yy @ preOrden zz)" | ||
− | value "preOrden (N e (N c (H a) (H d)) (N g (H f) (H h)))" | + | value "preOrden (N e (N c (H a) (H d)) (N g (H f) (H h))) |
− | -- "= [e,c, | + | = [e,c,a,d,g,f,h]" -- "True" |
+ | |||
+ | value "preOrden (N e (N c (N a1 (H a2) (H a3)) (H d)) (N g (H f) (H h))) | ||
+ | = [e, c, a1, a2, a3, d, g, f, h]" -- "True" | ||
text {* | text {* | ||
Línea 53: | Línea 58: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
fun postOrden :: "'a arbol ⇒ 'a list" where | fun postOrden :: "'a arbol ⇒ 'a list" where | ||
− | "postOrden | + | "postOrden (H x) = [x]" |
+ | |"postOrden (N x yy zz)= postOrden yy @ postOrden zz @ [x]" | ||
− | value "postOrden (N e (N c (H a) (H d)) (N g (H f) (H h))) | + | value "postOrden (N e (N c (H a) (H d)) (N g (H f) (H h))) |
− | + | = [a,d,c,f,h,g,e]" -- "True" | |
text {* | text {* | ||
Línea 70: | Línea 77: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
fun inOrden :: "'a arbol ⇒ 'a list" where | fun inOrden :: "'a arbol ⇒ 'a list" where | ||
− | "inOrden | + | "inOrden (H x) = [x]" |
+ | |"inOrden (N x yy zz) = inOrden yy @ [x] @ inOrden zz" | ||
− | value "inOrden (N e (N c (H a) (H d)) (N g (H f) (H h))) | + | value "inOrden (N e (N c (H a) (H d)) (N g (H f) (H h))) |
− | + | = [a,c,d,e,f,g,h]" -- "True" | |
text {* | text {* | ||
Línea 86: | Línea 95: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
fun espejo :: "'a arbol ⇒ 'a arbol" where | fun espejo :: "'a arbol ⇒ 'a arbol" where | ||
− | "espejo | + | "espejo (H x) = (H x)" |
+ | |"espejo (N x yy zz) = (N x (espejo zz) (espejo yy))" | ||
− | value "espejo (N e (N c (H a) (H d)) (N g (H f) (H h))) | + | value "espejo (N e (N c (H a) (H d)) (N g (H f) (H h))) |
− | + | = N e (N g (H h) (H f)) (N c (H d) (H a))" -- "True" | |
text {* | text {* | ||
Línea 99: | Línea 110: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
lemma "preOrden (espejo a) = rev (postOrden a)" | lemma "preOrden (espejo a) = rev (postOrden a)" | ||
− | + | by (induct a, simp_all) | |
text {* | text {* | ||
Línea 109: | Línea 121: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
lemma "postOrden (espejo a) = rev (preOrden a)" | lemma "postOrden (espejo a) = rev (preOrden a)" | ||
− | + | by (induct a, simp_all) | |
text {* | text {* | ||
Línea 119: | Línea 132: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
theorem "inOrden (espejo a) = rev (inOrden a)" | theorem "inOrden (espejo a) = rev (inOrden a)" | ||
− | + | by (induct a, simp_all) | |
text {* | text {* | ||
Línea 131: | Línea 145: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
fun raiz :: "'a arbol ⇒ 'a" where | fun raiz :: "'a arbol ⇒ 'a" where | ||
− | "raiz | + | "raiz (H x) = x" |
+ | |"raiz (N x _ _ ) = x" | ||
− | value "raiz (N e (N c (H a) (H d)) (N g (H f) (H h)))" -- " | + | value "raiz (N e (N c (H a) (H d)) (N g (H f) (H h))) = e" -- "True" |
text {* | text {* | ||
Línea 146: | Línea 162: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
fun extremo_izquierda :: "'a arbol ⇒ 'a" where | fun extremo_izquierda :: "'a arbol ⇒ 'a" where | ||
− | "extremo_izquierda | + | "extremo_izquierda (H x) = x" |
+ | |"extremo_izquierda (N _ yy _) = extremo_izquierda yy" | ||
value "extremo_izquierda (N e (N c (H a) (H d)) (N g (H f) (H h)))" -- "= a" | value "extremo_izquierda (N e (N c (H a) (H d)) (N g (H f) (H h)))" -- "= a" | ||
Línea 161: | Línea 179: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
fun extremo_derecha :: "'a arbol ⇒ 'a" where | fun extremo_derecha :: "'a arbol ⇒ 'a" where | ||
− | "extremo_derecha | + | "extremo_derecha (H x) = x" |
+ | |"extremo_derecha (N _ _ zz) = extremo_derecha zz" | ||
value "extremo_derecha (N e (N c (H a) (H d)) (N g (H f) (H h)))" -- "= h" | value "extremo_derecha (N e (N c (H a) (H d)) (N g (H f) (H h)))" -- "= h" | ||
Línea 173: | Línea 193: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
+ | (*jeshorcob: al intentar hacer la prueba automática, el sistema se | ||
+ | pregunta por el caso en el que la lista sea vacía. Debemos indicar que | ||
+ | esto no puede ocurrir y para ello añadimos un lema auxiliar.*) | ||
+ | lemma a1: "inOrden a ≠ []" | ||
+ | by (induct a, simp_all) | ||
+ | |||
+ | --"jeshorcob" | ||
theorem "last (inOrden a) = extremo_derecha a" | theorem "last (inOrden a) = extremo_derecha a" | ||
− | + | by (induct a, simp_all add: a1) | |
text {* | text {* | ||
Línea 183: | Línea 211: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
+ | (*jeshorcob: aquí necesitamos el mismo lema por un motivo similar*) | ||
theorem "hd (inOrden a) = extremo_izquierda a" | theorem "hd (inOrden a) = extremo_izquierda a" | ||
− | + | by (induct a, simp_all add: a1) | |
text {* | text {* | ||
Línea 193: | Línea 223: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
theorem "hd (preOrden a) = last (postOrden a)" | theorem "hd (preOrden a) = last (postOrden a)" | ||
− | + | by (induct a, simp_all) | |
text {* | text {* | ||
Línea 203: | Línea 234: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
theorem "hd (preOrden a) = raiz a" | theorem "hd (preOrden a) = raiz a" | ||
− | + | by (induct a, simp_all) | |
text {* | text {* | ||
Línea 213: | Línea 245: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
theorem "hd (inOrden a) = raiz a" | theorem "hd (inOrden a) = raiz a" | ||
+ | quickcheck | ||
oops | oops | ||
+ | (*Encuentra el contrajemplo: | ||
+ | a = N a⇩1 (H a⇩2) (H a⇩1) | ||
+ | |||
+ | ya que evaluando se ve que son distindos: | ||
+ | hd (inOrden a) = a⇩2 | ||
+ | raiz a = a⇩1 | ||
+ | *) | ||
text {* | text {* | ||
Línea 223: | Línea 264: | ||
*} | *} | ||
+ | --"jeshorcob" | ||
theorem "last (postOrden a) = raiz a" | theorem "last (postOrden a) = raiz a" | ||
− | + | by (induct a, simp_all) | |
end | end | ||
</source> | </source> |
Revisión del 15:03 29 nov 2014
header {* R7: Recorridos de árboles *}
theory R7
imports Main
begin
text {*
---------------------------------------------------------------------
Ejercicio 1. Definir el tipo de datos arbol para representar los
árboles binarios que tiene información en los nodos y en las hojas.
Por ejemplo, el árbol
e
/ \
/ \
c g
/ \ / \
a d f h
se representa por "N e (N c (H a) (H d)) (N g (H f) (H h))".
---------------------------------------------------------------------
*}
datatype 'a arbol = H "'a" | N "'a" "'a arbol" "'a arbol"
value "N e (N c (H a) (H d)) (N g (H f) (H h))"
text {*
---------------------------------------------------------------------
Ejercicio 2. 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 e (N c (H a) (H d)) (N g (H f) (H h)))
= [e,c,a,d,g,f,h]
---------------------------------------------------------------------
*}
--"jeshorcob"
fun preOrden :: "'a arbol ⇒ 'a list" where
"preOrden (H x) = [x]"
|"preOrden (N x yy zz) = x#(preOrden yy @ preOrden zz)"
value "preOrden (N e (N c (H a) (H d)) (N g (H f) (H h)))
= [e,c,a,d,g,f,h]" -- "True"
value "preOrden (N e (N c (N a1 (H a2) (H a3)) (H d)) (N g (H f) (H h)))
= [e, c, a1, a2, a3, d, g, f, h]" -- "True"
text {*
---------------------------------------------------------------------
Ejercicio 3. 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 e (N c (H a) (H d)) (N g (H f) (H h)))
= [a,d,c,f,h,g,e]
---------------------------------------------------------------------
*}
--"jeshorcob"
fun postOrden :: "'a arbol ⇒ 'a list" where
"postOrden (H x) = [x]"
|"postOrden (N x yy zz)= postOrden yy @ postOrden zz @ [x]"
value "postOrden (N e (N c (H a) (H d)) (N g (H f) (H h)))
= [a,d,c,f,h,g,e]" -- "True"
text {*
---------------------------------------------------------------------
Ejercicio 4. Definir la función
inOrden :: "'a arbol ⇒ 'a list"
tal que (inOrden a) es el recorrido in orden del árbol a. Por
ejemplo,
inOrden (N e (N c (H a) (H d)) (N g (H f) (H h)))
= [a,c,d,e,f,g,h]
---------------------------------------------------------------------
*}
--"jeshorcob"
fun inOrden :: "'a arbol ⇒ 'a list" where
"inOrden (H x) = [x]"
|"inOrden (N x yy zz) = inOrden yy @ [x] @ inOrden zz"
value "inOrden (N e (N c (H a) (H d)) (N g (H f) (H h)))
= [a,c,d,e,f,g,h]" -- "True"
text {*
---------------------------------------------------------------------
Ejercicio 5. Definir la función
espejo :: "'a arbol ⇒ 'a arbol"
tal que (espejo a) es la imagen especular del árbol a. Por ejemplo,
espejo (N e (N c (H a) (H d)) (N g (H f) (H h)))
= N e (N g (H h) (H f)) (N c (H d) (H a))
---------------------------------------------------------------------
*}
--"jeshorcob"
fun espejo :: "'a arbol ⇒ 'a arbol" where
"espejo (H x) = (H x)"
|"espejo (N x yy zz) = (N x (espejo zz) (espejo yy))"
value "espejo (N e (N c (H a) (H d)) (N g (H f) (H h)))
= N e (N g (H h) (H f)) (N c (H d) (H a))" -- "True"
text {*
---------------------------------------------------------------------
Ejercicio 6. Demostrar que
preOrden (espejo a) = rev (postOrden a)
---------------------------------------------------------------------
*}
--"jeshorcob"
lemma "preOrden (espejo a) = rev (postOrden a)"
by (induct a, simp_all)
text {*
---------------------------------------------------------------------
Ejercicio 7. Demostrar que
postOrden (espejo a) = rev (preOrden a)
---------------------------------------------------------------------
*}
--"jeshorcob"
lemma "postOrden (espejo a) = rev (preOrden a)"
by (induct a, simp_all)
text {*
---------------------------------------------------------------------
Ejercicio 8. Demostrar que
inOrden (espejo a) = rev (inOrden a)
---------------------------------------------------------------------
*}
--"jeshorcob"
theorem "inOrden (espejo a) = rev (inOrden a)"
by (induct a, simp_all)
text {*
---------------------------------------------------------------------
Ejercicio 9. Definir la función
raiz :: "'a arbol ⇒ 'a"
tal que (raiz a) es la raiz del árbol a. Por ejemplo,
raiz (N e (N c (H a) (H d)) (N g (H f) (H h))) = e
---------------------------------------------------------------------
*}
--"jeshorcob"
fun raiz :: "'a arbol ⇒ 'a" where
"raiz (H x) = x"
|"raiz (N x _ _ ) = x"
value "raiz (N e (N c (H a) (H d)) (N g (H f) (H h))) = e" -- "True"
text {*
---------------------------------------------------------------------
Ejercicio 10. Definir la función
extremo_izquierda :: "'a arbol ⇒ 'a"
tal que (extremo_izquierda a) es el nodo más a la izquierda del árbol
a. Por ejemplo,
extremo_izquierda (N e (N c (H a) (H d)) (N g (H f) (H h))) = a
---------------------------------------------------------------------
*}
--"jeshorcob"
fun extremo_izquierda :: "'a arbol ⇒ 'a" where
"extremo_izquierda (H x) = x"
|"extremo_izquierda (N _ yy _) = extremo_izquierda yy"
value "extremo_izquierda (N e (N c (H a) (H d)) (N g (H f) (H h)))" -- "= a"
text {*
---------------------------------------------------------------------
Ejercicio 11. Definir la función
extremo_derecha :: "'a arbol ⇒ 'a"
tal que (extremo_derecha a) es el nodo más a la derecha del árbol
a. Por ejemplo,
extremo_derecha (N e (N c (H a) (H d)) (N g (H f) (H h))) = h
---------------------------------------------------------------------
*}
--"jeshorcob"
fun extremo_derecha :: "'a arbol ⇒ 'a" where
"extremo_derecha (H x) = x"
|"extremo_derecha (N _ _ zz) = extremo_derecha zz"
value "extremo_derecha (N e (N c (H a) (H d)) (N g (H f) (H h)))" -- "= h"
text {*
---------------------------------------------------------------------
Ejercicio 12. Demostrar o refutar
last (inOrden a) = extremo_derecha a
---------------------------------------------------------------------
*}
--"jeshorcob"
(*jeshorcob: al intentar hacer la prueba automática, el sistema se
pregunta por el caso en el que la lista sea vacía. Debemos indicar que
esto no puede ocurrir y para ello añadimos un lema auxiliar.*)
lemma a1: "inOrden a ≠ []"
by (induct a, simp_all)
--"jeshorcob"
theorem "last (inOrden a) = extremo_derecha a"
by (induct a, simp_all add: a1)
text {*
---------------------------------------------------------------------
Ejercicio 13. Demostrar o refutar
hd (inOrden a) = extremo_izquierda a
---------------------------------------------------------------------
*}
--"jeshorcob"
(*jeshorcob: aquí necesitamos el mismo lema por un motivo similar*)
theorem "hd (inOrden a) = extremo_izquierda a"
by (induct a, simp_all add: a1)
text {*
---------------------------------------------------------------------
Ejercicio 14. Demostrar o refutar
hd (preOrden a) = last (postOrden a)
---------------------------------------------------------------------
*}
--"jeshorcob"
theorem "hd (preOrden a) = last (postOrden a)"
by (induct a, simp_all)
text {*
---------------------------------------------------------------------
Ejercicio 15. Demostrar o refutar
hd (preOrden a) = raiz a
---------------------------------------------------------------------
*}
--"jeshorcob"
theorem "hd (preOrden a) = raiz a"
by (induct a, simp_all)
text {*
---------------------------------------------------------------------
Ejercicio 16. Demostrar o refutar
hd (inOrden a) = raiz a
---------------------------------------------------------------------
*}
--"jeshorcob"
theorem "hd (inOrden a) = raiz a"
quickcheck
oops
(*Encuentra el contrajemplo:
a = N a⇩1 (H a⇩2) (H a⇩1)
ya que evaluando se ve que son distindos:
hd (inOrden a) = a⇩2
raiz a = a⇩1
*)
text {*
---------------------------------------------------------------------
Ejercicio 17. Demostrar o refutar
last (postOrden a) = raiz a
---------------------------------------------------------------------
*}
--"jeshorcob"
theorem "last (postOrden a) = raiz a"
by (induct a, simp_all)
end