Diferencia entre revisiones de «Relación 3»
De Razonamiento automático (2018-19)
|  (Página creada con «<source lang="isabelle"> chapter {* R3: Razonamiento sobre programas *}  theory R3_Razonamiento_sobre_programas imports Main  begin  text {* -------------------------------…») | |||
| Línea 23: | Línea 23: | ||
|    ------------------------------------------------------------------- *} |    ------------------------------------------------------------------- *} | ||
| − | + | (* "La demostración detallada es" *) | |
| + | |||
| + | (* benber *) | ||
| lemma "sumaImpares n = n*n" | 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 | ||
| text {* ---------------------------------------------------------------   | text {* ---------------------------------------------------------------   | ||
| Línea 46: | Línea 65: | ||
|    ------------------------------------------------------------------- *} |    ------------------------------------------------------------------- *} | ||
| + | (* benber *) | ||
| lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)" | 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 | ||
| text {* ---------------------------------------------------------------   | text {* ---------------------------------------------------------------   | ||
| Línea 79: | Línea 114: | ||
|    ------------------------------------------------------------------- *} |    ------------------------------------------------------------------- *} | ||
| + | (* benber *) | ||
| lemma "todos (λy. y=x) (copia n x)" | 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 | ||
| text {* ---------------------------------------------------------------   | text {* ---------------------------------------------------------------   | ||
| Línea 117: | Línea 168: | ||
|    "factI n = factI' n 1" |    "factI n = factI' n 1" | ||
| + | (* benber *) | ||
| lemma fact: "factI' n x = x * factR n" | 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 | ||
| text {* ---------------------------------------------------------------   | text {* ---------------------------------------------------------------   | ||
| Línea 125: | Línea 202: | ||
|    ------------------------------------------------------------------- *} |    ------------------------------------------------------------------- *} | ||
| + | (* benber *) | ||
| corollary "factI n = factR n" | 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 | ||
| text {* ---------------------------------------------------------------   | text {* ---------------------------------------------------------------   | ||
| Línea 145: | Línea 229: | ||
|    ------------------------------------------------------------------- *} |    ------------------------------------------------------------------- *} | ||
| + | (* benber *) | ||
| lemma "amplia xs y = xs @ [y]" | 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 | ||
| end | end | ||
| </source> | </source> | ||
Revisión del 20:02 23 nov 2018
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" *)
(* 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
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 *)
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
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. 
  ------------------------------------------------------------------- *}
(* benber *)
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
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
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
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 *)
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
end
