Diferencia entre revisiones de «Relación 3»
De Razonamiento automático (2017-18)
| (No se muestran 7 ediciones intermedias de 5 usuarios) | |||
| Línea 1: | Línea 1: | ||
| − | <source lang=" | + | <source lang="isabelle"> | 
| chapter {* R3: Razonamiento sobre programas *} | chapter {* R3: Razonamiento sobre programas *} | ||
| Línea 24: | Línea 24: | ||
| -- "La demostración detallada es" | -- "La demostración detallada es" | ||
| − | (*anddonram edupalhid rafcabgon luicedval jescudero macmerflo diwu2 rafferrod*) | + | (* anddonram edupalhid rafcabgon luicedval jescudero macmerflo diwu2 | 
| + |    rafferrod cesgongut jospermon1 davperriv *)   | ||
| lemma "sumaImpares n = n*n" | lemma "sumaImpares n = n*n" | ||
| proof (induct n) | proof (induct n) | ||
| Línea 31: | Línea 32: | ||
|    fix n |    fix n | ||
|    assume HI: "sumaImpares n=n*n" |    assume HI: "sumaImpares n=n*n" | ||
| − |    have "sumaImpares (Suc n) =sumaImpares n + (2*n+1)" by simp | + |    have "sumaImpares (Suc n) = sumaImpares n + (2*n+1)" by simp | 
|    also have "... = n*n + (2*n+1)" using HI by simp |    also have "... = n*n + (2*n+1)" using HI by simp | ||
|    finally show "sumaImpares (Suc n) =(Suc n) *(Suc n)" by simp |    finally show "sumaImpares (Suc n) =(Suc n) *(Suc n)" by simp | ||
| qed | qed | ||
| − | |||
| text {* ---------------------------------------------------------------   | text {* ---------------------------------------------------------------   | ||
| Línea 55: | Línea 55: | ||
|       sumaPotenciasDeDosMasUno n = 2^(n+1) |       sumaPotenciasDeDosMasUno n = 2^(n+1) | ||
|    ------------------------------------------------------------------- *} |    ------------------------------------------------------------------- *} | ||
| − | (*anddonram edupalhid luicedval macmerflo jescudero diwu2 rafferrod*) | + | |
| + | (* anddonram edupalhid luicedval macmerflo jescudero diwu2 rafferrod jospermon1 davperriv *) | ||
| lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)" | lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)" | ||
| proof (induct n) | proof (induct n) | ||
| Línea 62: | Línea 63: | ||
|    fix n |    fix n | ||
|    assume HI: "sumaPotenciasDeDosMasUno n = 2^(n+1)" |    assume HI: "sumaPotenciasDeDosMasUno n = 2^(n+1)" | ||
| − |    have "sumaPotenciasDeDosMasUno (Suc n) = sumaPotenciasDeDosMasUno n + 2^(n+1)" by simp | + |    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^(n+1) +  2^(n+1)" using HI by simp | ||
|    also have "... = 2^((n+1)+1)" by simp |    also have "... = 2^((n+1)+1)" by simp | ||
| − |    finally show "sumaPotenciasDeDosMasUno (Suc n) = 2^(Suc n + 1)" by simp | + |    finally show "sumaPotenciasDeDosMasUno (Suc n) = 2^(Suc n + 1)"   | 
| + |     by simp | ||
|    qed |    qed | ||
| − | (*rafcabgon*) | + | (* rafcabgon cesgongut *) | 
| lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)" | lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)" | ||
| proof (induct n) | proof (induct n) | ||
| Línea 75: | Línea 78: | ||
|    fix n |    fix n | ||
|    assume HI: "sumaPotenciasDeDosMasUno n = 2^(n+1)" |    assume HI: "sumaPotenciasDeDosMasUno n = 2^(n+1)" | ||
| − |    have "sumaPotenciasDeDosMasUno (Suc n) = sumaPotenciasDeDosMasUno n + 2^(n+1)" by simp | + |    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^(n+1) + 2^(n+1)" using HI by simp | ||
|    also have "... = 2^((Suc n) + 1)" by simp |    also have "... = 2^((Suc n) + 1)" by simp | ||
| − |    finally show "sumaPotenciasDeDosMasUno (Suc n) = 2 ^ (Suc n + 1)" by simp | + |    finally show "sumaPotenciasDeDosMasUno (Suc n) = 2 ^ (Suc n + 1)"   | 
| + |     by simp | ||
| qed | qed | ||
| Línea 111: | Línea 116: | ||
|    ------------------------------------------------------------------- *} |    ------------------------------------------------------------------- *} | ||
| − | (*anddonram edupalhid rafcabgon luicedval macmerflo jescudero diwu2*) | + | (* anddonram edupalhid rafcabgon luicedval macmerflo jescudero diwu2 jospermon1 davperriv cesgongut *) | 
| lemma "todos (λy. y=x) (copia n x)" | lemma "todos (λy. y=x) (copia n x)" | ||
| proof (induct n) | proof (induct n) | ||
| Línea 118: | Línea 123: | ||
|    fix n |    fix n | ||
|    assume HI: "todos (λy. y=x) (copia n x)" |    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 | + |    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 "... =  todos (λy. y=x) (copia n x) " by simp | ||
|    show "todos (λy. y=x) (copia (Suc n) x)"  using HI by simp |    show "todos (λy. y=x) (copia (Suc n) x)"  using HI by simp | ||
| qed | qed | ||
| − | (*rafferrod*) | + | (* rafferrod *) | 
| lemma "todos (λy. y=x) (copia n x)" | lemma "todos (λy. y=x) (copia n x)" | ||
| proof (induct n) | proof (induct n) | ||
| Línea 130: | Línea 136: | ||
|    fix n |    fix n | ||
|    assume HI: "todos (λy. y=x) (copia n x)" |    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 | + |    have "todos (λy. y=x) (copia (Suc n) x) =   | 
| − |    also have "... = (((λy. y=x) x) ∧ (todos (λy. y=x) (copia n x)))" by simp | + |         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 "... = True" using HI by simp |    also have "... = True" using HI by simp | ||
|    finally show "todos (λy. y=x) (copia (Suc n) x)" by simp |    finally show "todos (λy. y=x) (copia (Suc n) x)" by simp | ||
| Línea 171: | Línea 179: | ||
|    "factI n = factI' n 1" |    "factI n = factI' n 1" | ||
| − | (*anddonram rafferrod*) | + | (* anddonram rafferrod *) | 
| lemma fact: "factI' n x = x * factR n" | lemma fact: "factI' n x = x * factR n" | ||
| proof (induct n arbitrary: x) | proof (induct n arbitrary: x) | ||
| Línea 179: | Línea 187: | ||
|    assume HI: "⋀x. factI' n x = x * factR n"   |    assume HI: "⋀x. factI' n x = x * factR n"   | ||
|    fix x |    fix x | ||
| − |    have "factI' (Suc n) x = factI' n (x*Suc n)" by simp | + |    have "factI' (Suc n) x = factI' n (x * Suc n)" by simp | 
|    also have "... = (x * Suc n) * factR n" using HI by simp   |    also have "... = (x * Suc n) * factR n" using HI by simp   | ||
|    also have "... = x * (Suc n * factR n)" by (simp del:mult_Suc) |    also have "... = x * (Suc n * factR n)" by (simp del:mult_Suc) | ||
| Línea 185: | Línea 193: | ||
| qed | qed | ||
| − | (* edupalhid rafcabgon luicedval macmerflo diwu2 jescudero*) | + | (* edupalhid rafcabgon luicedval macmerflo diwu2 jescudero jospermon1 davperriv  *) | 
| lemma fact': "factI' n x = x * factR n" | lemma fact': "factI' n x = x * factR n" | ||
| proof (induct n arbitrary: x) | proof (induct n arbitrary: x) | ||
| Línea 207: | Línea 215: | ||
|       factI n = factR n |       factI n = factR n | ||
|    ------------------------------------------------------------------- *} |    ------------------------------------------------------------------- *} | ||
| − | (*anddonram edupalhid rafcabgon luicedval macmerflo diwu2 jescudero rafferrod*) | + | |
| + | (* anddonram edupalhid rafcabgon luicedval macmerflo diwu2 jescudero   | ||
| + |    rafferrod jospermon1 davperriv *) | ||
| corollary "factI n = factR n" | corollary "factI n = factR n" | ||
| proof - | proof - | ||
| Línea 234: | Línea 244: | ||
| − | (*anddonram rafcabgon luicedval macmerflo diwu2 jescudero*) | + | (* anddonram rafcabgon luicedval macmerflo diwu2 jescudero rafferrod jospermon1 davperriv cesgongut *) | 
| lemma "amplia xs y = xs @ [y]" | lemma "amplia xs y = xs @ [y]" | ||
| proof (induct xs) | proof (induct xs) | ||
Revisión actual del 20:40 14 jul 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"
(* anddonram edupalhid rafcabgon luicedval jescudero macmerflo diwu2
   rafferrod cesgongut jospermon1 davperriv *) 
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
  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)
  ------------------------------------------------------------------- *}
(* anddonram edupalhid luicedval macmerflo jescudero diwu2 rafferrod jospermon1 davperriv *)
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^((n+1)+1)" by simp
  finally show "sumaPotenciasDeDosMasUno (Suc n) = 2^(Suc n + 1)" 
    by simp
  qed
(* rafcabgon cesgongut *)
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
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. 
  ------------------------------------------------------------------- *}
(* anddonram edupalhid rafcabgon luicedval macmerflo jescudero diwu2 jospermon1 davperriv cesgongut *)
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
  show "todos (λy. y=x) (copia (Suc n) x)"  using HI by simp
qed
(* rafferrod *)
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)))" 
    by simp
  also have "... = True" using HI by simp
  finally show "todos (λy. y=x) (copia (Suc n) x)" 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"
(* anddonram rafferrod *)
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" 
  fix x
  have "factI' (Suc n) x = factI' n (x * Suc n)" by simp
  also have "... = (x * Suc n) * factR n" using HI by simp 
  also have "... = x * (Suc n * factR n)" by (simp del:mult_Suc)
  finally show "factI' (Suc n) x = x * factR (Suc n)" by simp
qed
(* edupalhid rafcabgon luicedval macmerflo diwu2 jescudero jospermon1 davperriv  *)
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 (x * Suc n)"   by simp
    also have "... = (x * Suc n) * factR n" using HI by simp
    also have "... = x * (Suc n * factR n)" by (simp del: mult_Suc)
    also have "... = x * factR (Suc n)" by simp
    finally show "factI' (Suc n) x = x * factR (Suc n)" by simp
  qed
qed
text {* --------------------------------------------------------------- 
  Ejercicio 4.3. Escribir la demostración detallada de
     factI n = factR n
  ------------------------------------------------------------------- *}
(* anddonram edupalhid rafcabgon luicedval macmerflo diwu2 jescudero 
   rafferrod jospermon1 davperriv *)
corollary "factI n = factR n"
proof -
  fix n
  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]
  ------------------------------------------------------------------- *}
(* anddonram rafcabgon luicedval macmerflo diwu2 jescudero rafferrod jospermon1 davperriv cesgongut *)
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
  also have "... =(a # xs) @[y] " using HI by simp
  finally show "amplia (a#xs) y= (a # xs) @[y]" by simp
qed
(* edupalhid *)
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
