Diferencia entre revisiones de «Relación 3»
De Razonamiento automático (2013-14)
m (Protegió «Relación 3» ([edit=sysop] (indefinido) [move=sysop] (indefinido))) |
|
(Sin diferencias)
|
Revisión del 09:56 7 dic 2013
header {* R3: Razonamiento sobre programas en Isabelle/HOL *}
theory R3
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
------------------------------------------------------------------ *}
-- "maresccas4 irealetei diecabmen1 domlloriv pabflomar marescpla"
fun sumaImpares :: "nat ⇒ nat" where
"sumaImpares 0 = 0"
| "sumaImpares (Suc n) = 2 * n + 1 + sumaImpares n"
value "sumaImpares 5" -- "= 25"
-- "julrobrel"
fun sumaImpares2 :: "nat ⇒ nat" where
"sumaImpares2 0 = 0"
| "sumaImpares2 (Suc n) = (2 * (Suc n) - 1) + sumaImpares2 n"
value "sumaImpares2 5" -- "= 25"
text {* ---------------------------------------------------------------
Ejercicio 2. Demostrar que
sumaImpares n = n*n
------------------------------------------------------------------- *}
-- "maresccas4 pabflomar juaruipav marescpla"
lemma "sumaImpares n = n*n"
proof (induct n)
show "sumaImpares 0 = 0*0" by simp
next
fix n
assume HI: "sumaImpares n = n*n"
have "sumaImpares (Suc n) = 2 * n + 1 + sumaImpares n" by (simp only: sumaImpares.simps(2))
also have "... = 2 * n + 1 + n * n" using HI by simp
finally show "sumaImpares (Suc n) = (Suc n) * (Suc n)" by simp
qed
-- "irealetei diecabmen1 domlloriv"
lemma "sumaImpares n = n*n"
proof (induct n)
show "sumaImpares (0) = 0 * 0" by simp
next
fix n
assume HI: "sumaImpares n = n * n"
have "sumaImpares (Suc n) = (2*n + 1) + sumaImpares n " by simp
also have "... = 2*n + 1 + (n * n)" using HI by simp
also have "...= (n+1)*(n+1)" by simp
finally show "sumaImpares (Suc n) =Suc n * Suc n" by simp
qed
-- "Sobrarían varios parentesis."
-- "julrobrel"
lemma "sumaImpares n = n*n"
proof (induct n)
show "sumaImpares 0 = 0 * 0" by simp
next
fix n
assume HI: "sumaImpares n = n * n"
have "sumaImpares (Suc n) = (2 * (Suc n) - 1) + sumaImpares n" by simp
also have "...= (n+1)*(n+1)" using HI by simp
finally show "sumaImpares (Suc n) = (Suc n) * (Suc n)" by simp
qed
-- "ElyIvan"
lemma "sumaImpares n = n*n"
proof -
by (induc n)
show "sumaImpares 0 = 0*0" by simp
next
assume HI: "sumaImpares n = n * n"
have "sumaImpares n = 2 * n - 1) + sumaImpares ( n- 1 )" by sumanImpares2
also have "...= 2 * n - 1 + (n - 1) * ( n - 1)" using HI by simp
also have "...= 2 * n -1 + n * n - 2 * n + 1" by simp
also have "...= n * n" by simp
finally show "sumaImpares n = n * 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
------------------------------------------------------------------ *}
-- "maresccas4 irealetei diecabmen1 domlloriv pabflomar julrobrel"
fun sumaPotenciasDeDosMasUno :: "nat ⇒ nat" where
"sumaPotenciasDeDosMasUno 0 = 2"
| "sumaPotenciasDeDosMasUno (Suc n) = 2^(Suc n) + sumaPotenciasDeDosMasUno n"
value "sumaPotenciasDeDosMasUno 3" -- "= 16"
-- "ElyIvan"
fun sumaPotenciasDeDosMasUno2 :: "nat ⇒ nat" where
"sumaPotenciasDeDosMasUno2 0 = 2"
| "sumaPotenciasDeDosMasUno2 n = 2^n + sumaPotenciasDeDosMasUno2 (n-1)"
value "sumaPotenciasDeDosMasUno2 1" -- "= 16"
text {* ---------------------------------------------------------------
Ejercicio 4. Demostrar que
sumaPotenciasDeDosMasUno n = 2^(n+1)
------------------------------------------------------------------- *}
-- "maresccas4 pabflomar juaruipav"
lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)"
proof (induct n)
show "sumaPotenciasDeDosMasUno 0 = 2^(0+1)" by simp
next
fix n
assume HI: "sumaPotenciasDeDosMasUno n = 2^(n+1)"
have "sumaPotenciasDeDosMasUno (Suc n) = 2^(Suc n) + sumaPotenciasDeDosMasUno n" by (simp only: sumaPotenciasDeDosMasUno.simps(2))
also have "... = 2^(Suc n) + 2^(n+1)" using HI by simp
also have "... = 2^(n+1) + 2^(n+1)" by simp
also have "... = 2 * 2^(n+1)" by simp
also have "... = 2 * 2^(Suc n)" by simp
finally show "sumaPotenciasDeDosMasUno (Suc n) = 2^((Suc n)+1)" by simp
qed
--"irealetei"
lemma "sumaPotenciasDeDosMasUno n = 2^(n + 1)"
proof (induct n)
show "sumaPotenciasDeDosMasUno (0) = 2^(0 + 1)" by simp
next
fix n
assume HI: "sumaPotenciasDeDosMasUno n = 2^(n + 1)"
have "sumaPotenciasDeDosMasUno (n+1) = 2^(n+1) + sumaPotenciasDeDosMasUno(n)" by simp
also have "... = 2^(n+1) + 2^(n+1)" using HI by simp
also have "...=2*2^(n+1)" by simp
also have "...=2^(n+2)" by simp
finally show "sumaPotenciasDeDosMasUno (Suc n) = 2^(Suc n +1)" by simp
qed
--"diecabmen1, marescpla"
lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)"
proof (induct n)
show "sumaPotenciasDeDosMasUno 0 = 2^(0+1)" by simp
next
fix n
assume HI: "sumaPotenciasDeDosMasUno n = 2^(n+1)"
have "sumaPotenciasDeDosMasUno (Suc n) = 2^(Suc n)+sumaPotenciasDeDosMasUno n" by simp
also have "… = (2^(Suc n)) + (2^(n+1))" using HI by simp
also have "… = 2^(n+1) + 2^(n+1)" by simp
also have "… = 2 * 2^(Suc n)" by simp
finally show "sumaPotenciasDeDosMasUno (Suc n) = 2^(Suc n +1)" by simp
qed
--"en el primer have yo especifiqué la regla usada: by (simp only: sumaPotenciasDeDosMasUno.simps(2)) marescpla"
--"domlloriv"
lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)"
proof (induct n)
show "sumaPotenciasDeDosMasUno 0 = 2^(0+1)" by simp
next
fix n
assume HI: "sumaPotenciasDeDosMasUno n = 2^(n+1)"
have "sumaPotenciasDeDosMasUno (Suc n) = 2^(Suc n) + sumaPotenciasDeDosMasUno n" by simp
also have "... = 2^(Suc n) + 2^(n+1)" using HI by simp
finally show "sumaPotenciasDeDosMasUno (Suc n) = 2^((Suc n)+1)" by simp
qed
-- "julrobrel"
lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)"
proof (induct n)
show "sumaPotenciasDeDosMasUno 0 = 2^(0+1)" by simp
next
fix n
assume HI: "sumaPotenciasDeDosMasUno n = 2^(n+1)"
have "sumaPotenciasDeDosMasUno (Suc n) = 2^(Suc n) + sumaPotenciasDeDosMasUno n" by simp
also have "...= 2^((n+1)+1)" using HI by simp
finally show "sumaPotenciasDeDosMasUno (Suc n) = 2^((Suc n)+1)" by simp
qed
-- "ElyIvan"
lemma "sumaPotenciasDeDosMasUno2 n = 2^(n+1)"
proof (induct n)
show "sumaPotenciasDeDosMasUno2 0 = 2^(0+1)" by simp
next
fix n
assume HI "sumaPotenciasDeDosMasUno2 n = 2^(n+1)"
have "sumaPotenciasDeDosMasUno2 n = 2^n + sumaPotenciasDeDosMasUno2 (n-1)" by sumaPotenciasDeDosMasUno2 2
also have "...= 2^n + 2 ^((n-1)+1)" by HI
also have "...= 2^n + 2^n" by simp
finally show "sumaPotenciasDeDosMasUno2 n = 2^(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]
------------------------------------------------------------------ *}
-- "maresccas4 irealetei diecabmen1 domlloriv pabflomar julrobrel marescpla"
fun copia :: "nat ⇒ 'a ⇒ 'a list" where
"copia 0 x = []"
| "copia (Suc n) x = x # copia n x"
value "copia 3 x" -- "= [x,x,x]"
-- "ElyIvan"
fun copia2 :: "nat ⇒ 'a ⇒ 'a list" where
"copia2 0 x = []"
| "copia2 n x = x # copia2 ( n - 1 ) x"
value "copia2 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 ∧
----------------------------------------------------------------- *}
-- "maresccas4 irealetei diecabmen1 domlloriv pabflomar julrobrel marescpla"
fun todos :: "('a ⇒ bool) ⇒ 'a list ⇒ bool" where
"todos p [] = True"
| "todos p (x#xs) = (p x ∧ todos 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.
------------------------------------------------------------------- *}
-- "maresccas4"
lemma "todos (λy. y=x) (copia n x)"
proof (induct n)
show "todos (λy. y=x) (copia 0 x)" by simp
next
fix n
assume HI: "todos (λy. y=x) (copia n x)"
have "todos (λy. y=x) (copia (Suc n) x) = todos (λy. y=x) (x # copia n x)" by (simp only: copia.simps(2))
also have "... = todos (λy. y=x) (copia n x)" by simp
finally show "todos (λy. y=x) (copia (Suc n) x)" using HI by simp
qed
-- "irealetei diecabmen1 pabflomar"
lemma "todos (λy. y=x) (copia n x)"
proof (induct n)
show "todos (λy. y = x) (copia 0 x)" by simp
next
fix n
assume HI: " todos (λy. y = x) (copia n x)"
have " todos (λy. y = x) (copia (Suc n) x)= todos (λy. y = x) (x#copia n x)" by simp
also have "... = todos (λy. y = x) (copia n x)" by simp
also have "... = True" using HI by simp
finally show "todos (λy. y = x) (copia (Suc n) x)" by simp
qed
-- "domlloriv"
lemma "todos (λy. y=x) (copia n x)"
proof (induct n)
show "todos (λy. y=x) (copia 0 x)" by simp
next
fix n
assume HI:"todos (λy. y=x) (copia n x)"
have "todos (λy. y=x) (copia (Suc n) x) = ((λy. y=x) x ∧ todos (λy. y=x) (copia n x))" using HI by simp
also have "... = True" using HI by simp
finally show "todos (λy. y=x) (copia (Suc n) x)" by simp
qed
-- "julrobrel"
lemma "todos (λy. y=x) (copia n x)"
proof (induct n)
show "todos (λy. y=x) (copia 0 x)" by simp
next
fix n
assume HI: "todos (λy. y = x) (copia n x)"
have "todos (λy. y = x) (copia (Suc n) x)=todos (λy. y = x) (x # copia n x)" by simp
also have "...=(λy. y = x) x ∧ todos (λy. y = x) (copia n x)" using HI by simp
finally show "todos (λy. y = x) (copia (Suc n) x)" by simp
qed
-- "juaruipav"
lemma "todos (λy. y=x) (copia n x)"
proof (induct n)
show "todos (λy. y=x) (copia 0 x)" by simp
next
fix n
assume HI: "todos (λy. y=x) (copia n x)"
have "todos (λy. y=x) (copia (Suc n) x)= todos (λy. y=x)(x# (copia n x))" by simp
also have "...=todos (λy. y=x)[x]" using HI by simp
finally show "todos (λy. y=x) (copia (Suc n) x)" by simp
qed
--"marescpla"
lemma "todos (λy. y=x) (copia n x)"
proof(induct n)
show "todos (λy. y=x) (copia 0 x)" by (simp only: copia.simps(1) todos.simps(1))
next
fix n
assume HI : "todos (λy. y=x) (copia n x)"
have "todos (λy. y=x) (copia (Suc n) x) = ((λy. y=x) (hd (copia (Suc n) x)) ∧ todos (λy. y=x) (tl(copia (Suc n) x)))" by simp
also have "...= ((λy. y=x) (hd (x # copia n x))∧ todos (λy. y=x) (tl(copia (Suc n) x)))" by (simp only: copia.simps(2))
also have " ...= ((λy. y=x) x ∧ todos (λy. y=x) (tl(copia (Suc n) x)))" by (simp only: hd.simps)
also have "...=(todos (λy. y=x) (tl(copia (Suc n) x)))" by simp
also have "...=(todos (λy. y=x) (tl(x # copia n x)))" by (simp only: copia.simps(2))
also have "...=(todos (λy. y=x) (copia n x))" by (simp only: tl.simps)
also have "...= True" using HI by simp
finally show "todos (λy. y=x) (copia (Suc n) x)" by simp
qed
--"jajaja ya he visto que me he complicado demasiado. Es mejor resolverlo en el otro orden"
text {* ---------------------------------------------------------------
Ejercicio 8. Definir la función
factR :: nat ⇒ nat
tal que (factR n) es el factorial de n. Por ejemplo,
factR 4 = 24
------------------------------------------------------------------ *}
-- "maresccas4 irealetei diecabmen1 domlloriv pabflomar julrobrel marescpla"
fun factR :: "nat ⇒ nat" where
"factR 0 = 1"
| "factR (Suc n) = Suc n * factR 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"
-- "maresccas4"
lemma fact: "factI' n x = x * factR n"
proof (induct n arbitrary: x)
show "⋀x. factI' 0 x = x * factR 0" by simp
next
fix n
assume HI: "⋀x. factI' n x = x * factR n"
show "⋀x. factI' (Suc n) x = x * factR (Suc n)"
proof -
fix x
have "factI' (Suc n) x = factI' n (Suc n)*x" by (simp only: factI'.simps(2))
also have "... = x * factI' n (Suc n)" by simp
also have "... = x * ((Suc n) * factR n)" using HI by simp
finally show "factI' (Suc n) x = x * factR (Suc n)" by simp
qed
qed
-- "irealetei" (*Éste ha sido la muerte a pellizcos, aún leyéndome todo el tema.*)
lemma fact2: "factI' n x = x * factR n"
proof (induct n arbitrary: x)
show "⋀x. factI' 0 x = x * factR 0" by simp
next
fix n x
assume HI: "⋀x. factI' n x = x * factR n"
have "factI' (Suc n) x = factI' n (Suc n)*x" by simp
also have "... = x * factI' n (Suc n)" by simp
also have "... = x * (Suc n * factR n)" using HI by simp
also have "... = x * factR (Suc n)" by simp
finally show "factI' (Suc n) x = x * factR (Suc n)" by simp
qed
-- "diecabmen1" (*Es el mismo que de irealetei solo que me parece que se puede eliminar el último also have y ponerlo directo en el finally.*)
lemma fact3: "factI' n x = x * factR n"
proof (induct n arbitrary: x)
show "⋀x. factI' 0 x = x * factR 0" by simp
next
fix n x
assume HI: "⋀x. factI' n x = x * factR n"
have "factI' (Suc n) x = factI' n (Suc n)*x" by simp
also have "… = x * ((Suc n) * factR n)" using HI by simp
finally show "factI' (Suc n) x = x * factR (Suc n)" by simp
qed
-- "domlloriv"
lemma fact4: "factI' n x = x * factR n"
proof (induct n arbitrary: x)
show "⋀x. factI' 0 x = x * factR 0" by simp (*Para todo x se cumple*)
next
fix n
assume HI: "⋀x. factI' n x = x * factR n"
have "factI' (Suc n) x = factI' n (Suc n)*x" using HI by simp
also have "... = x * (Suc n * factR n)" using HI by simp
finally show "⋀x. factI' (Suc n) x = x * factR (Suc n)" using HI by simp
qed
(* Comentario: "using HI" se necesita sólo una vez. *)
-- "julrobrel"
lemma fact5: "factI' n x = x * factR n"
proof (induct n arbitrary: x)
show "⋀x. factI' 0 x = x * factR 0" by simp
next
fix n x
assume HI: "⋀x. factI' n x = x * factR n"
have "factI' (Suc n) x = factI' n (Suc n)*x" by simp
also have "...=x * ((Suc n) * factR n)" using HI by simp
finally show "factI' (Suc n) x = x* factR (Suc n)" by simp
qed
-- "juaruipav"
lemma fact6: "factI' n x = x * factR n"
proof (induct n arbitrary: x)
show "⋀x. factI' 0 x= x*factR 0" by simp
next
fix n
assume HI: "⋀x. factI' n x = x * factR n"
have "factI'(Suc n) x= factI' n (Suc n)*x" by simp
also have "...=x*((Suc n)*factR n)" using HI by simp
finally show "⋀x. factI' (Suc n) x = x*factR (Suc n)" by simp
qed
-- " juaruipav: No me termina de funcionar"
-- "pabflomar los entiendo casi todos pero yo no habría sacado ni uno de ellos"
--"marescpla (no me resuelve el paso de inducción"
lemma fact: "factI' n x = x * factR n"
proof (induct n)
show "factI' 0 x = x* factR 0" by (simp only: factI'.simps(1) factR.simps(1))
next
fix n
assume HI : "factI' n x = x * factR n"
have " factI' (Suc n) x= factI' n (Suc n)*x" by (simp only: factI'.simps(2))
also have "...= (Suc n)*x * factR n" using HI by simp
also have " ... = x * (Suc n)*factR n" by simp
also have " ... = x * factR (Suc n)" by (simp only: factR.simps(2))
finally show " factI' (Suc n) x = x * factR (Suc n)" by simp
qed
text {* ---------------------------------------------------------------
Ejercicio 10. Demostrar que
factI n = factR n
------------------------------------------------------------------- *}
-- "maresccas4 julrobrel"
corollary "factI n = factR n"
proof -
show "factI n = factR n" by (simp add:fact)
qed
-- "irealetei diecabmen1 domlloriv pabflomar juaruipav"
corollary "factI n = factR n"
by (simp add: fact)
--"marescpla"
corollary "factI n = factR n"
proof-
have "factI n = factI' n 1" by (simp only: factI.simps(1))
also have " ...= 1*factR n" by (simp add: fact)
finally show "factI n = factR n" by simp
qed
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]
------------------------------------------------------------------ *}
-- "maresccas4 irealetei diecabmen1 domlloriv julrobrel marescpla"
fun amplia :: "'a list ⇒ 'a ⇒ 'a list" where
"amplia [] y = [y]"
| "amplia (x#xs) y = x # (amplia xs y)"
value "amplia [d,a] t" -- "= [d,a,t]"
text {* ---------------------------------------------------------------
Ejercicio 12. Demostrar que
amplia xs y = xs @ [y]
------------------------------------------------------------------- *}
-- "maresccas4 marescpla"
lemma "amplia xs y = xs @ [y]"
proof (induct xs)
show "amplia [] y = [] @ [y]" by simp
next
fix x xs
assume HI: "amplia xs y = xs @ [y]"
have "amplia (x#xs) y = x # (amplia xs y)" by (simp only: amplia.simps(2))
also have "... = x # xs @ [y]" using HI by simp
finally show "amplia (x#xs) y = (x#xs) @ [y]" by simp
qed
--"irealetei diecabmen1 domlloriv pabflomar juaruipav"
lemma "amplia xs y = xs @ [y]"
proof (induct xs)
show "amplia [] y = [] @ [y]" by simp
next
fix a xs
assume HI: "amplia xs y = xs @ [y]"
have "amplia (a#xs) y = a # (amplia xs y)" by simp
also have "... = (a # xs) @ [y]" using HI by simp
finally show "amplia (a#xs) y = (a # xs) @ [y]" by simp
qed
-- "julrobrel"
lemma "amplia xs y = xs @ [y]"
proof (induct xs)
show "amplia [] y = [] @ [y]" by simp
next
fix x xs
assume HI:"amplia xs y = xs @ [y]"
have "amplia (x#xs) y = x # (amplia xs y)" by simp
also have "...= x # (xs @ [y])" using HI by simp
finally show "amplia (x#xs) y = (x#xs) @ [y]" by simp
qed
end