Acciones

Diferencia entre revisiones de «Relación 3»

De Razonamiento automático (2016-17)

m
Línea 261: Línea 261:
 
qed
 
qed
  
 +
(*danrodcha*)
 
lemma "amplia xs y = xs @ [y]" (is "?P xs")
 
lemma "amplia xs y = xs @ [y]" (is "?P xs")
 
proof (induct xs)
 
proof (induct xs)

Revisión del 02:02 12 nov 2016

chapter {* R3: Razonamiento sobre programas *}

theory R3_Razonamiento_sobre_programas
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
  ------------------------------------------------------------------- *}

-- "La demostración detallada es"
(* crigomgom fraortmoy marpoldia1 ivamenjim *)
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) = sumaImpares n + 2*n + 1" by simp
  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

(*danrodcha ; es la misma demostración que la anterior pero uso ?P para sustituir la propiedad.*)
lemma "sumaImpares n = n*n" (is "?P n")
proof (induct n)
  show "?P 0" by simp
next
  fix n
  assume HI: "?P n"
  have "sumaImpares (Suc n) = sumaImpares n + 2*n + 1" by simp
  also have "... = n*n + 2*n + 1" using HI by simp
  also have "... = Suc n * Suc n" by simp 
  finally show "?P (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)
  ------------------------------------------------------------------- *}

(*crigomgom ivamenjim danrodcha*)
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) = sumaPotenciasDeDosMasUno n + 2^(n+1)" by simp
  also have "... = 2 ^ (n + 1) + 2 ^ (n + 1)" using HI by simp
  also have "... = 2 ^ ((Suc n) + 1)" by simp
  finally show "sumaPotenciasDeDosMasUno (Suc n) = 2 ^ (Suc n + 1)" by simp
qed

(* fraortmoy *)
(* es la misma demostración, pero quise probar a delimitar lo que se usa en el "by simp" *)

lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)"
proof (induct n)
  show "sumaPotenciasDeDosMasUno 0 = 2 ^ (0 + 1)" by simp
next
  fix n
  assume H1:" 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 H1 by simp
  also have "... = 2 ^ (Suc n + 1)" by simp
  finally show "sumaPotenciasDeDosMasUno (Suc n) =  2 ^ (Suc n + 1)" by simp
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. 
  ------------------------------------------------------------------- *}

(* crigomgom *)
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 "... = ((x = x) ∧ (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


(* fraortmoy *)
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 "... = True" by (simp only: todos.simps(1))
  show "todos (λy. y = x) (copia 0 x)" by simp
next
 fix n 
 assume H1 : " todos (λy. y = x) (copia n x) "
 have "todos (λy. y = x) (copia (Suc n) x) = ((todos (λy. y = x) (x#[])) ∧  (todos (λy. y = x) (copia n x) )) " by simp
 also have " ... = (todos (λy. y = x) (x#[]))" using H1 by simp
 also have " ... = ((λy. y = x) x ∧ todos (λy. y = x) [])" by simp
 also have " ... = True" by simp
 finally show " todos (λy. y = x) (copia (Suc n) x)" by simp
qed

(*danrodcha*)
lemma "todos (λy. y=x) (copia n x)" (is "?P n")
proof (induct n)
  show "?P 0" by simp
next
  fix n
  assume HI: "?P n"
  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
  also have "... = todos (λy. y=x) (copia n x)" by simp
  finally show "?P (Suc n)" using HI by simp
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"

(* fraortmoy , danrodcha*)
lemma fact: "factI' n x = x * factR n"
proof (induct n)
  show "factI' 0 x = x * factR 0" by simp
next
  fix n
  assume H1 :  "factI' n x = x * 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" using H1 by simp (* no entiendo por qué no hace esto bien y luego todo funciona *)
  also have "... = x * factR (Suc n)" by (simp del: mult_Suc)
  finally show "factI' (Suc n) x = x * factR (Suc n)" by simp
qed

text {* --------------------------------------------------------------- 
  Ejercicio 4.3. Escribir la demostración detallada de
     factI n = factR n
  ------------------------------------------------------------------- *}

(* fraortmoy danrodcha *)
corollary "factI n = factR n"
proof -
  have "factI n = factI' n 1" by simp
  also have "... = 1 * factR n" by (simp add: fact)
  finally show "factI n = factR n" by simp
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]
  ------------------------------------------------------------------- *}

(* crigomgom fraortmoy *)
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
  also have "... = (x # xs) @ [y]" by simp
  finally show "amplia (x # xs) y = (x # xs) @ [y]" by simp
qed

(*danrodcha*)
lemma "amplia xs y = xs @ [y]" (is "?P xs")
proof (induct xs)
  show "?P []" by simp
next
  fix x xs
  assume HI: "?P xs"
  have "amplia (x # xs) y = x # amplia xs y" by simp
  also have "... = (x # xs) @ [y]" using HI by simp
  finally show "?P (x#xs)" by simp
qed

end