header {* R4: Razonamiento sobre programas en Isabelle/HOL *}
theory R4
imports Main
begin
text {* ---------------------------------------------------------------
Ejercicio 1. Definir la función
sumaImpares :: nat ⇒ nat
tal que (sumaImpares n) es la suma de los n primeros números
impares. Por ejemplo,
sumaImpares 5 = 25
------------------------------------------------------------------ *}
(*Solución M.Cumplido*)
fun sumaImpares :: "nat ⇒ nat" where
"sumaImpares 0 = 0"
| "sumaImpares (Suc n) = (2 * (Suc n) - 1) + sumaImpares n"
value "sumaImpares 5" -- "= 25"
fun sumaImpares_lecc :: "nat ⇒ nat" where
"sumaImpares_lecc 0 = 0"
| "sumaImpares_lecc (Suc n) = 2 * n + 1 + sumaImpares_lecc n"
text {* ---------------------------------------------------------------
Ejercicio 2. Demostrar que
sumaImpares n = n*n
------------------------------------------------------------------- *}
lemma "sumaImpares n = n*n"
oops
lemma "sumaImpares_lecc n = n*n"
proof (induct n)
show "sumaImpares_lecc 0 = 0*0" by simp
next
fix n
assume HI: "sumaImpares_lecc n = n*n"
have "sumaImpares_lecc (Suc n) = 2 * n + 1 + sumaImpares_lecc n" by simp
also have "... = 2 * n + 1 + n * n" using HI by simp
also have "... = (Suc n)*(Suc n)" by simp
finally show "sumaImpares_lecc (Suc n) = (Suc n)*(Suc n)" by simp
qed
text {* ---------------------------------------------------------------
Ejercicio 3. Definir la función
sumaPotenciasDeDosMasUno :: nat ⇒ nat
tal que
(sumaPotenciasDeDosMasUno n) = 1 + 2^0 + 2^1 + 2^2 + ... + 2^n.
Por ejemplo,
sumaPotenciasDeDosMasUno 3 = 16
------------------------------------------------------------------ *}
fun sumaPotenciasDeDosMasUno :: "nat ⇒ nat" where
"sumaPotenciasDeDosMasUno n = undefined"
fun sumaPotenciasDeDosMasUno_lecc :: "nat ⇒ nat" where
"sumaPotenciasDeDosMasUno_lecc 0 = 1 + 2^0"
| "sumaPotenciasDeDosMasUno_lecc (Suc n) = 2 ^ (n + 1) + sumaPotenciasDeDosMasUno_lecc n"
value "sumaPotenciasDeDosMasUno 3" -- "= 16"
text {* ---------------------------------------------------------------
Ejercicio 4. Demostrar que
sumaPotenciasDeDosMasUno n = 2^(n+1)
------------------------------------------------------------------- *}
lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)"
oops
lemma "sumaPotenciasDeDosMasUno_lecc n = 2^(n+1)"
proof (induct n)
have "sumaPotenciasDeDosMasUno_lecc 0 = 1 + 2^0" by simp
also have "... = 2^(0+1)" by simp
finally show "sumaPotenciasDeDosMasUno_lecc 0 = 2^(0+1)" by simp
next
fix n::nat
assume HI: "sumaPotenciasDeDosMasUno_lecc n = 2^(n+1)"
have "sumaPotenciasDeDosMasUno_lecc (Suc n) = 2^(n+1)+ sumaPotenciasDeDosMasUno_lecc n" by simp
also have "... = 2^(n+1)+2^(n+1)" using HI by simp
also have "... = 2^(n+2)" by simp
also have "... = 2^(Suc n + 1)" by simp
finally show "sumaPotenciasDeDosMasUno_lecc (Suc n) = 2^(Suc n + 1)" by simp
qed
text {* ---------------------------------------------------------------
Ejercicio 5. Definir la función
copia :: nat ⇒ 'a ⇒ 'a list
tal que (copia n x) es la lista formado por n copias del elemento
x. Por ejemplo,
copia 3 x = [x,x,x]
------------------------------------------------------------------ *}
fun copia :: "nat ⇒ 'a ⇒ 'a list" where
"copia n x = undefined"
fun copia_lecc :: "nat ⇒ 'a ⇒ 'a list" where
"copia_lecc 0 x = []"
| "copia_lecc (Suc n) x = (x#copia_lecc n x)"
value "copia 3 x" -- "= [x,x,x]"
text {* ---------------------------------------------------------------
Ejercicio 6. Definir la función
todos :: ('a ⇒ bool) ⇒ 'a list ⇒ bool
tal que (todos p xs) se verifica si todos los elementos de xs cumplen
la propiedad p. Por ejemplo,
todos (λx. x>(1::nat)) [2,6,4] = True
todos (λx. x>(2::nat)) [2,6,4] = False
Nota: La conjunción se representa por ∧
----------------------------------------------------------------- *}
fun todos :: "('a ⇒ bool) ⇒ 'a list ⇒ bool" where
"todos p xs = undefined"
fun todos_lecc :: "('a ⇒ bool) ⇒ 'a list ⇒ bool" where
"todos_lecc p [] = True"
| "todos_lecc p (x#xs) = (p x ∧ todos_lecc p xs)"
value "todos (λx. x>(1::nat)) [2,6,4]" -- "= True"
value "todos (λx. x>(2::nat)) [2,6,4]" -- "= False"
text {* ---------------------------------------------------------------
Ejercicio 7. Demostrar que todos los elementos de (copia n x) son
iguales a x.
------------------------------------------------------------------- *}
lemma "todos (λy. y=x) (copia n x)"
oops
lemma "todos_lecc (λy. y=x) (copia_lecc n x)"
proof (induct n)
have "todos_lecc (λy. y=x) (copia_lecc 0 x) = todos_lecc (λy. y=x) []" by simp
also have "... = True" by simp
finally show "todos_lecc (λy. y=x) (copia_lecc 0 x)" by simp
next
fix n::nat
assume HI: "todos_lecc (λy. y=x) (copia_lecc n x)"
have "todos_lecc (λy. y=x) (copia_lecc (Suc n) x) = todos_lecc (λy. y=x) (x#copia_lecc n x)" by simp
also have "... = ((λy. y=x) x ∧ todos_lecc (λy. y=x) (copia_lecc n x))" by simp
also have "... = (True ∧ todos_lecc (λy. y=x) (copia_lecc n x))" by simp
also have "... = (True ∧ True)" using HI by simp
also have "... = True" by simp
finally show "todos_lecc (λy. y=x) (copia_lecc (Suc n) x)" by simp
qed
text {* ---------------------------------------------------------------
Ejercicio 8. Definir la función
factR :: nat ⇒ nat
tal que (factR n) es el factorial de n. Por ejemplo,
factR 4 = 24
------------------------------------------------------------------ *}
fun factR :: "nat ⇒ nat" where
"factR n = undefined"
fun factR_lecc :: "nat ⇒ nat" where
"factR_lecc 0 = 1"
| "factR_lecc (Suc n) = (n+1) * factR_lecc n"
value "factR 4" -- "= 24"
text {* ---------------------------------------------------------------
Ejercicio 9. Se considera la siguiente definición iterativa de la
función factorial
factI :: "nat ⇒ nat" where
factI n = factI' n 1
factI' :: nat ⇒ nat ⇒ nat" where
factI' 0 x = x
factI' (Suc n) x = factI' n (Suc n)*x
Demostrar que, para todo n y todo x, se tiene
factI' n x = x * factR n
------------------------------------------------------------------- *}
fun factI' :: "nat ⇒ nat ⇒ nat" where
"factI' 0 x = x"
| "factI' (Suc n) x = factI' n (Suc n)*x"
fun factI :: "nat ⇒ nat" where
"factI n = factI' n 1"
value "factI 4" -- "= 24"
lemma fact: "factI' n x = x * factR n"
oops
text {* ---------------------------------------------------------------
Ejercicio 10. Demostrar que
factI n = factR n
------------------------------------------------------------------- *}
corollary "factI n = factR n"
oops
text {* ---------------------------------------------------------------
Ejercicio 11. Definir, recursivamente y sin usar (@), la función
amplia :: 'a list ⇒ 'a ⇒ 'a list
tal que (amplia xs y) es la lista obtenida añadiendo el elemento y al
final de la lista xs. Por ejemplo,
amplia [d,a] t = [d,a,t]
------------------------------------------------------------------ *}
fun amplia :: "'a list ⇒ 'a ⇒ 'a list" where
"amplia xs y = undefined"
fun amplia_lecc :: "'a list ⇒ 'a ⇒ 'a list" where
"amplia_lecc [] y = [y]"
| "amplia_lecc (x#xs) y = x#(amplia_lecc xs y)"
value "amplia [d,a] t" -- "= [d,a,t]"
text {* ---------------------------------------------------------------
Ejercicio 12. Demostrar que
amplia xs y = xs @ [y]
------------------------------------------------------------------- *}
lemma "amplia xs y = xs @ [y]"
oops
lemma "amplia_lecc xs y = xs @ [y]"
proof (induct xs)
have 1:"amplia_lecc [] y = [y]" by simp
have "[] @ [y] = [y]" by simp
hence "[y] = [] @ [y]" ..
with 1 show "amplia_lecc [] y = [] @ [y]" by simp
next
fix x xs
assume HI: "amplia_lecc xs y = xs @ [y]"
have "amplia_lecc (x#xs) y = x#(amplia_lecc xs y)" by simp
also have "... = x#(xs @ [y])" using HI by simp
also have "... = (x#xs) @ [y]"
proof (induct xs)
show "x#([]@[y]) = [x]@[y]" by simp
next
fix a as
assume HI: "x#(as @ [y]) = (x#as)@[y]"
have "x#((a#as)@[y]) = x#(a#(as @ [y]))" by simp
also have "... = x#([a] @ as @ [y])" by simp
also have "... = [x] @ [a] @ as @ [y]" by simp
also have "... = (x # a # as) @ [y]" by simp
finally show "x#((a#as)@[y]) = (x # a # as) @ [y]" by simp
qed
finally show "amplia_lecc (x#xs) y = (x#xs) @ [y]" by simp
qed
end