Relación 3
De Razonamiento automático (2018-19)
Revisión del 12:05 10 ene 2019 de Jalonso (discusión | contribuciones)
chapter {* R3: Razonamiento sobre programas *}
theory R3_Razonamiento_sobre_programas_alu
imports Main
begin
text {* ---------------------------------------------------------------
Ejercicio 1.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
------------------------------------------------------------------ *}
fun sumaImpares :: "nat ⇒ nat" where
"sumaImpares 0 = 0"
| "sumaImpares (Suc n) = sumaImpares n + (2*n+1)"
text {* ---------------------------------------------------------------
Ejercicio 1.2. Escribir la demostración detallada de
sumaImpares n = n*n
------------------------------------------------------------------- *}
(* benber *)
lemma "sumaImpares n = n*n"
proof (induct n)
show "sumaImpares 0 = 0*0" by (simp only: sumaImpares.simps(1))
next
fix n
assume HI: "sumaImpares n = n*n"
have "sumaImpares (Suc n) = sumaImpares n + (2*n+1)"
by (simp only: sumaImpares.simps(2))
also have "... = n*n + (2*n+1)" by (simp only: HI)
also have "... = n*n + (n+n+1)" by (simp only:)
also have "... = n*n + n + n + 1" by (simp only:)
also have "... = n*n + n*1 + n + 1" by (simp only:)
also have "... = n*(n+1) + n + 1"
by (simp only: Nat.add_mult_distrib2)
also have "... = n*(n+1) + (n+1)" by (simp only:)
also have "... = (n+1)*n + (n+1)"
by (simp only: Groups.ab_semigroup_mult_class.mult.commute)
also have "... = (n+1)*n + (n+1)*1" by (simp only: Nat.nat_mult_1_right)
also have "... = (n+1)*(n+1)" by (simp only: Nat.add_mult_distrib2)
also have "... = (Suc n) * (Suc n)" by (simp only: Nat.Suc_eq_plus1)
finally show "sumaImpares (Suc n) = (Suc n) * (Suc n)" .
qed
(* manperjim pabalagon josgomrom4 hugrubsan enrparalv juacanrod *)
lemma "sumaImpares n = n*n"
proof (induct n)
show "sumaImpares 0 = 0*0" by (simp only: sumaImpares.simps(1))
next
fix n
assume HI: "sumaImpares n = n*n"
have "sumaImpares (Suc n) = sumaImpares n + 2 * n + 1"
by (simp only: sumaImpares.simps(2))
also have "... = n*n + 2*n + 1" using HI by simp
also have "... = n*n + n + n + 1" by (simp only: Nat.add_mult_distrib)
also have "... = (n+1)*n + (n + 1)" by (simp only: Nat.add_mult_distrib)
also have "... = (n+1)*n + (n + 1)*1" by (simp only: Nat.nat_mult_1_right)
also have "... = (n+1)*(n+1)" by (simp only: Nat.add_mult_distrib2)
finally show "sumaImpares (Suc n) = (Suc n)*(Suc n)"
by (simp only: Nat.Suc_eq_plus1)
qed
(* marfruman1 raffergon2 aribatval alfmarcua cammonagu gleherlop alikan
pabbergue giafus1 *)
lemma "sumaImpares n = n*n"
proof (induct n)
show "sumaImpares 0 = 0*0" by (simp only: sumaImpares.simps(1))
next
fix n
assume HI: "sumaImpares n = n*n"
have "sumaImpares (Suc n) = sumaImpares n + 2*n +1"
by (simp only: sumaImpares.simps(2))
also have "... = n*n + 2*n + 1" using HI by simp
also have "... = (Suc n)*(Suc n)" by simp
finally show "sumaImpares (Suc n) = (Suc n)*(Suc n)" by simp
qed
text {* ---------------------------------------------------------------
Ejercicio 2.1. 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 0 = 2"
| "sumaPotenciasDeDosMasUno (Suc n) =
sumaPotenciasDeDosMasUno n + 2^(n+1)"
text {* ---------------------------------------------------------------
Ejercicio 2.2. Escribir la demostración detallada de
sumaPotenciasDeDosMasUno n = 2^(n+1)
------------------------------------------------------------------- *}
(* benber josgomrom4 *)
lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)"
proof (induct n)
have "sumaPotenciasDeDosMasUno 0 = 2"
by (simp only: sumaPotenciasDeDosMasUno.simps(1))
also have "... = 2^(0+1)" by simp
finally show "sumaPotenciasDeDosMasUno 0 = 2^(0+1)" .
next
fix n
assume HI: "sumaPotenciasDeDosMasUno n = 2^(n+1)"
have "sumaPotenciasDeDosMasUno (Suc n) =
sumaPotenciasDeDosMasUno n + 2^(n+1)"
by (simp only: sumaPotenciasDeDosMasUno.simps(2))
also have "... = 2^(n+1) + 2^(n+1)" by (simp only: HI)
also have "... = 2 * 2^(n+1)" by (simp only:)
also have "... = 2 * 2^(Suc n)" by (simp only: Nat.Suc_eq_plus1)
also have "... = 2^(Suc (Suc n))"
by (simp only: Power.power_class.power.power_Suc)
also have "... = 2^((Suc n)+1)" by (simp only: Nat.Suc_eq_plus1)
finally show "sumaPotenciasDeDosMasUno (Suc n) = 2^((Suc n)+1)" .
qed
(* marfruman1 alfmarcua hugrubsan enrparalv gleherlop juacanrod *)
lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)"
proof (induct n)
have "sumaPotenciasDeDosMasUno 0 = 2"
by (simp only: sumaPotenciasDeDosMasUno.simps(1))
also have "... = 2^(0+1)" by simp
finally show "sumaPotenciasDeDosMasUno 0 = 2^(0+1)" by simp
next
fix n
assume HI: "sumaPotenciasDeDosMasUno n = 2^(n+1)"
have "sumaPotenciasDeDosMasUno (Suc n) =
sumaPotenciasDeDosMasUno n + 2^(n+1)"
by (simp only: sumaPotenciasDeDosMasUno.simps(2))
also have "... = 2^(n+1)+2^(n+1)" using HI by simp
also have "... =2*2^(n+1)" by simp
also have "... = 2^((Suc n)+1)" by simp
finally show "sumaPotenciasDeDosMasUno (Suc n) = 2^((Suc n)+1)"
by simp
qed
(* manperjim pabalagon raffergon2 aribatval cammonagu hugrubsan
pabbergue giafus1 *)
lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)"
proof (induct n)
have "sumaPotenciasDeDosMasUno 0 = 2"
by (simp only: sumaPotenciasDeDosMasUno.simps(1))
also have "... = 2^1" by (simp only: Power.power_one_right)
finally show "sumaPotenciasDeDosMasUno 0 = 2^(0+1)"
by (simp only: Nat.add_0)
next
fix n
assume HI: "sumaPotenciasDeDosMasUno n = 2^(n+1)"
have "sumaPotenciasDeDosMasUno (Suc n) =
sumaPotenciasDeDosMasUno n + 2^(n+1)"
by (simp only: sumaPotenciasDeDosMasUno.simps(2))
also have "... = 2^(n+1) + 2^(n+1)" by (simp only: HI)
also have "... = 2*2^(n+1)" by (simp only: Nat.add_mult_distrib)
also have "... = 2*2^(Suc n)" by (simp only: Nat.Suc_eq_plus1)
also have "... = 2^(Suc n)*2^1" by (simp only: Power.power_one_right)
finally show "sumaPotenciasDeDosMasUno (Suc n) = 2^(Suc n + 1)"
by (simp only: Power.power_add)
qed
text {* ---------------------------------------------------------------
Ejercicio 3.1. 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 0 x = []"
| "copia (Suc n) x = x # copia n x"
text {* ---------------------------------------------------------------
Ejercicio 3.2. 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
------------------------------------------------------------------ *}
fun todos :: "('a ⇒ bool) ⇒ 'a list ⇒ bool" where
"todos p [] = True"
| "todos p (x#xs) = (p x ∧ todos p xs)"
text {* ---------------------------------------------------------------
Ejercicio 3.2. Demostrar detalladamente que todos los elementos de
(copia n x) son iguales a x.
------------------------------------------------------------------- *}
(* manperjim benber raffergon2 aribatval josgomrom4 hugrubsan cammonagu
gleherlop marfruman1 pabbergue giafus1 juacanrod *)
lemma "todos (λy. y=x) (copia n x)"
proof (induct n)
fix x :: 'a
have "todos (λy. y=x) (copia 0 x) = todos (λy. y=x) []"
by (simp only: copia.simps(1))
also have "..." by (simp only: todos.simps(1))
finally show "todos (λy. y = x) (copia 0 x)" .
next
fix x :: 'a
fix n :: nat
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 "... = (((λy. y = x) x ) ∧ todos (λy. y = x) (copia n x) )"
by (simp only: todos.simps(2))
also have "..." by (simp only: HI)
finally show "todos (λy. y = x) (copia (Suc n) x)" .
qed
(* pabalagon *)
lemma "todos (λy. y=x) (copia n x)"
proof (induction n)
fix x:: 'a
have "todos (λy. y=x) (copia 0 x) = todos (λy. y=x) []" by simp
also have "..." by simp
finally show "todos (λy. y=x) (copia 0 x)" .
next
fix x:: 'a
fix n:: nat
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))" by simp
finally show "todos (λy. y=x) (copia (Suc n) x)" using HI by simp
qed
(* alfmarcua *)
lemma "todos (λy. y=x) (copia n x)"
proof (induct n)
have "todos (λy. y = x) (copia 0 x) = todos (λy. y = x) []"
by (simp only: copia.simps(1))
also have "..." by (simp only: todos.simps(1))
finally show "todos (λy. y = x) (copia 0 x)" .
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 "... = ( ( (λy. y = x) x ) ∧ todos (λy. y = x) (copia n x) )"
by (simp only: todos.simps(2))
also have "..." by (simp only: HI)
finally show "todos (λy. y = x) (copia (Suc n) x)" .
qed
text {* ---------------------------------------------------------------
Ejercicio 4.1. 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 0 = 1"
| "factR (Suc n) = Suc n * factR n"
text {* ---------------------------------------------------------------
Ejercicio 4.2. 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
Indicación: La propiedad mult_Suc es
(Suc m) * n = n + m * n
Puede que se necesite desactivarla en un paso con
(simp del: mult_Suc)
------------------------------------------------------------------- *}
fun factI' :: "nat ⇒ nat ⇒ nat" where
"factI' 0 x = x"
| "factI' (Suc n) x = factI' n (x * Suc n)"
fun factI :: "nat ⇒ nat" where
"factI n = factI' n 1"
(* benber *)
lemma fact: "factI' n x = x * factR n"
proof -
(* Reformulación para aplicar la hipótesis inductiva con un valor
distinto de x. *)
have "∀x. factI' n x = x * factR n"
proof (induct n)
show "∀x. factI' 0 x = x * factR 0"
proof
fix x
have "factI' 0 x = x" by (simp only: factI'.simps(1))
also have "... = x * 1" by (simp only:)
also have "... = x * factR 0" by (simp only: factR.simps(1))
finally show "factI' 0 x = x * factR 0" .
qed
next
show "⋀n. ∀x. factI' n x = x * factR n ⟹
∀x. factI' (Suc n) x = x * factR (Suc n)"
proof
fix x
fix n
assume HI: "∀y. factI' n y = y * factR n"
have "factI' (Suc n) x = factI' n (x * Suc n)"
by (simp only: factI'.simps(2))
also have "... = (x * Suc n) * factR n" by (simp only: HI)
also have "... = x * factR (Suc n)" by (simp only: factR.simps(2))
finally show "factI' (Suc n) x = x * factR (Suc n)" .
qed
qed
thus "factI' n x = x * factR n" by simp
qed
(* manperjim pabalagon raffergon2 aribatval josgomrom4 alfmarcua
hugrubsan cammonagu marfruman1 gleherlop pabbergue giafus1 juacanrod *)
lemma fact2: "factI' n x = x * factR n"
proof (induction n arbitrary: x)
fix x
have "factI' 0 x = x" by simp
also have "... = x * 1" by simp
finally show "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 (x * Suc n)"
by (simp only: factI'.simps(2))
also have "... = x * Suc n * factR n" using HI by simp
finally show "factI' (Suc n) x = x*factR (Suc n)"
by (simp only: factR.simps(2))
qed
qed
text {* ---------------------------------------------------------------
Ejercicio 4.3. Escribir la demostración detallada de
factI n = factR n
------------------------------------------------------------------- *}
(* benber *)
corollary "factI n = factR n"
proof -
fix n
have "factI n = factI' n 1" by (simp only: factI.simps(1))
also have "... = 1 * factR n" by (simp only: fact)
also have "... = factR n" by (simp only: Groups.mult_1)
finally show "factI n = factR n" .
qed
(* manperjim pabalagon raffergon2 josgomrom4 alfmarcua hugrubsan
marfruman1 gleherlop pabbergue giafus1 juacanrod *)
corollary "factI n = factR n"
proof -
have "factI n = factI' n 1" by (simp only: factI.simps(1))
also have "... = 1*factR n" using fact by simp
finally show "factI n = factR n" by (simp only: mult_1)
qed
text {* ---------------------------------------------------------------
Ejercicio 5.1. 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 [] y = [y]"
| "amplia (x#xs) y = x # amplia xs y"
text {* ---------------------------------------------------------------
Ejercicio 5.2. Escribir la demostración detallada de
amplia xs y = xs @ [y]
------------------------------------------------------------------- *}
(* benber aribatval alfmarcua giafus1 *)
lemma "amplia xs y = xs @ [y]"
proof (induct xs)
have "amplia [] y = [y]" by (simp only: amplia.simps(1))
also have "... = [] @ [y]" by (simp only: List.append.left_neutral)
finally show "amplia [] y = [] @ [y]" .
next
fix x
fix 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])" by (simp only: HI)
also have "... = (x#xs) @ [y]" by (simp only: List.append.append_Cons)
finally show "amplia (x#xs) y = (x#xs) @ [y]" .
qed
(* marfruman1 raffergon2 enrparalv alikan gleherlop juacanrod *)
lemma "amplia xs y = xs @ [y]"
proof (induct xs)
have "amplia [] y = [y]" by (simp only: amplia.simps(1))
also have "... = [] @ [y]" by simp
finally 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
also have "... = (x#xs) @ [y]" by simp
finally show "amplia (x#xs) y =(x#xs) @ [y]" by simp
qed
(* manperjim pabalagon josgomrom4 hugrubsan cammonagu pabbergue *)
lemma "amplia xs y = xs @ [y]"
proof (induction xs)
have "amplia [] y = [y]" by (simp only: amplia.simps(1))
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 only: amplia.simps(2))
also have "... = a # (xs @ [y])" using HI by simp
finally show "amplia (a#xs) y = (a#xs) @ [y]" by simp
qed
end