Diferencia entre revisiones de «Relación 3»
De Razonamiento automático (2016-17)
m (Texto reemplazado: «isar» por «isabelle») |
|||
(No se muestran 53 ediciones intermedias de 17 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" | ||
− | (* crigomgom fraortmoy marpoldia1 ivamenjim serrodcal rubgonmar ferrenseg juacabsou wilmorort *) | + | (* crigomgom fraortmoy marpoldia1 ivamenjim serrodcal rubgonmar |
+ | ferrenseg juacabsou wilmorort josgarsan lucnovdos paupeddeg bowma | ||
+ | pabrodmac dancorgar antsancab1 *) | ||
lemma "sumaImpares n = n*n" | lemma "sumaImpares n = n*n" | ||
proof (induct n) | proof (induct n) | ||
Línea 37: | Línea 39: | ||
qed | qed | ||
− | (*danrodcha | + | (* danrodcha: Es la misma demostración que la anterior pero uso ?P para |
+ | sustituir la propiedad.*) | ||
lemma "sumaImpares n = n*n" (is "?P n") | lemma "sumaImpares n = n*n" (is "?P n") | ||
proof (induct n) | proof (induct n) | ||
Línea 63: | Línea 66: | ||
qed | qed | ||
− | (* fracorjim1 | + | (* fracorjim1: Especifico la regla de simplificación y el paso del |
+ | desarrollo al cuadrado para hacerlo más legible. *) | ||
lemma "sumaImpares n = n*n" | lemma "sumaImpares n = n*n" | ||
proof (induct n) | proof (induct n) | ||
Línea 70: | Línea 74: | ||
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 only: sumaImpares.simps(2)) | + | 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 + 2 * n + 1" using HI by simp | ||
also have "... = Suc n * Suc n" by simp | also have "... = Suc n * Suc n" by simp | ||
Línea 76: | Línea 81: | ||
qed | qed | ||
− | (*palucoto anaprarod*) | + | (* palucoto anaprarod marcarmor13 jeamacpov *) |
lemma "sumaImpares n = n*n" | lemma "sumaImpares n = n*n" | ||
proof (induct n) | proof (induct n) | ||
Línea 87: | Línea 92: | ||
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 108: | Línea 112: | ||
------------------------------------------------------------------- *} | ------------------------------------------------------------------- *} | ||
− | (*crigomgom ivamenjim danrodcha serrodcal rubgonmar ferrenseg juacabsou wilmorort anaprarod marpoldia1*) | + | (*crigomgom ivamenjim danrodcha serrodcal rubgonmar ferrenseg juacabsou |
+ | wilmorort anaprarod marpoldia1 josgarsan paupeddeg pabrodmac bowma | ||
+ | dancorgar antsancab1 *) | ||
lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)" | lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)" | ||
proof (induct n) | proof (induct n) | ||
Línea 115: | Línea 121: | ||
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 | ||
− | (* fraortmoy | + | (* 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)" | lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)" | ||
proof (induct n) | proof (induct n) | ||
Línea 130: | Línea 137: | ||
fix n | fix n | ||
assume H1:" sumaPotenciasDeDosMasUno n = 2 ^ (n + 1)" | assume H1:" sumaPotenciasDeDosMasUno n = 2 ^ (n + 1)" | ||
− | have "sumaPotenciasDeDosMasUno (Suc n) = sumaPotenciasDeDosMasUno n + 2^(n + 1)" by (simp only : sumaPotenciasDeDosMasUno.simps(2)) | + | 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 ^ (n + 1) + 2 ^ (n + 1)" using H1 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 143: | Línea 153: | ||
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)" | + | have "sumaPotenciasDeDosMasUno (Suc n) = |
− | + | sumaPotenciasDeDosMasUno n + 2^(n+1)" | |
− | also have "sumaPotenciasDeDosMasUno (Suc n) = 2^(n+1)+2^(n+1)" using HI by simp | + | by (simp only: sumaPotenciasDeDosMasUno.simps(2)) |
− | finally show "sumaPotenciasDeDosMasUno (Suc n) = 2^((Suc n)+1)" by simp | + | also have "sumaPotenciasDeDosMasUno (Suc n) = 2^(n+1)+2^(n+1)" |
+ | using HI by simp | ||
+ | finally show "sumaPotenciasDeDosMasUno (Suc n) = 2^((Suc n)+1)" | ||
+ | by simp | ||
qed | qed | ||
− | (*pablucoto*) | + | (* pablucoto marcarmor13 lucnovdos jeamacpov *) |
lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)" | lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)" | ||
proof (induct n) | proof (induct n) | ||
Línea 156: | Línea 169: | ||
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 | ||
− | finally show " sumaPotenciasDeDosMasUno (Suc n) = 2 ^ (Suc n + 1)" by simp | + | finally show " sumaPotenciasDeDosMasUno (Suc n) = 2 ^ (Suc n + 1)" |
+ | by simp | ||
qed | qed | ||
− | (* fracorjim1 - Hago explícita toda la manipulación algebraica. Quizás excesivo.*) | + | (* fracorjim1 - Hago explícita toda la manipulación algebraica. Quizás |
+ | excesivo.*) | ||
lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)" | lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)" | ||
proof (induction n) | proof (induction n) | ||
Línea 168: | Línea 184: | ||
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)" | + | have "sumaPotenciasDeDosMasUno (Suc n) = |
+ | sumaPotenciasDeDosMasUno n + 2 ^ (n + 1)" | ||
by (simp only:sumaPotenciasDeDosMasUno.simps(2)) | by (simp only:sumaPotenciasDeDosMasUno.simps(2)) | ||
also have "... = 2 ^ (n + 1) + 2 ^ (n + 1)" using HI by simp | also have "... = 2 ^ (n + 1) + 2 ^ (n + 1)" using HI by simp | ||
Línea 174: | Línea 191: | ||
also have "... = 2 ^ ((n + 1) + 1)" by simp | also have "... = 2 ^ ((n + 1) + 1)" 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 207: | Línea 225: | ||
------------------------------------------------------------------- *} | ------------------------------------------------------------------- *} | ||
− | (* crigomgom ivamenjim serrodcal ferrenseg wilmorort *) | + | (* crigomgom ivamenjim serrodcal ferrenseg wilmorort juacabsou josgarsan |
+ | lucnovdos *) | ||
lemma "todos (λy. y=x) (copia n x)" | lemma "todos (λy. y=x) (copia n x)" | ||
proof (induct n) | proof (induct n) | ||
Línea 214: | Línea 233: | ||
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 "... = ((x = x) ∧ (todos (λy. y = 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 | 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 | ||
qed | qed | ||
− | |||
(* fraortmoy *) | (* fraortmoy *) | ||
lemma "todos (λy. y=x) (copia n x)" | lemma "todos (λy. y=x) (copia n x)" | ||
proof (induct n) | proof (induct n) | ||
− | have "todos (λy. y = x) (copia 0 x) = todos (λy. y = x) []" by (simp only: copia.simps(1)) | + | 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)) | also have "... = True" by (simp only: todos.simps(1)) | ||
show "todos (λy. y = x) (copia 0 x)" by simp | show "todos (λy. y = x) (copia 0 x)" by simp | ||
Línea 230: | Línea 250: | ||
fix n | fix n | ||
assume H1 : " todos (λy. y = x) (copia n x) " | assume H1 : " todos (λy. y = x) (copia n x) " | ||
− | have "todos (λy. y = x) (copia (Suc n) x) = ((todos (λy. y = x) (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 " ... = (todos (λy. y = x) (x#[]))" using H1 by simp | ||
also have " ... = ((λy. y = x) x ∧ todos (λy. y = x) [])" by simp | also have " ... = ((λy. y = x) x ∧ todos (λy. y = x) [])" by simp | ||
Línea 237: | Línea 259: | ||
qed | qed | ||
− | (*danrodcha rubgonmar*) | + | (* danrodcha rubgonmar *) |
lemma "todos (λy. y=x) (copia n x)" (is "?P n") | lemma "todos (λy. y=x) (copia n x)" (is "?P n") | ||
proof (induct n) | proof (induct n) | ||
Línea 244: | Línea 266: | ||
fix n | fix n | ||
assume HI: "?P n" | assume HI: "?P n" | ||
− | 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 "... = ((λy. y=x) x ∧ todos (λy. y=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 | also have "... = todos (λy. y=x) (copia n x)" by simp | ||
Línea 253: | Línea 276: | ||
lemma "todos (λy. y=x) (copia n x)" | lemma "todos (λy. y=x) (copia n x)" | ||
proof (induct n) | proof (induct n) | ||
− | have "todos (λy. y=x) (copia 0 x) = todos (λy. y=x) []" by (simp only: copia.simps(1)) | + | have "todos (λy. y=x) (copia 0 x) = todos (λy. y=x) []" |
+ | by (simp only: copia.simps(1)) | ||
also have "todos (λy. y=x) []" by (simp only: todos.simps(1)) | also have "todos (λy. y=x) []" by (simp only: todos.simps(1)) | ||
show "todos (λy. y=x) (copia 0 x)" by simp | show "todos (λy. y=x) (copia 0 x)" by simp | ||
Línea 259: | Línea 283: | ||
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))" | + | have "todos (λy. y=x) (copia (Suc n) x) = |
+ | todos (λy. y=x) (x # (copia n x))" | ||
by (simp only: copia.simps(2)) | by (simp only: copia.simps(2)) | ||
− | also have "todos (λy. y=x) (x#(copia n x)) = ((λy. y=x) x ∧ todos (λy. y=x) (copia n x))" | + | also have "todos (λy. y=x) (x#(copia n x)) = |
− | + | ((λy. y=x) x ∧ todos (λy. y=x) (copia n x))" | |
− | also have "todos (λy. y=x) (copia (Suc n) x) = ((λy. y=x) x)" using HI by simp | + | by (simp only: todos.simps(2)) |
+ | also have "todos (λy. y=x) (copia (Suc n) x) = ((λy. y=x) x)" | ||
+ | using HI by simp | ||
also have "((λy. y=x) x) = True" by simp | also have "((λy. y=x) x) = True" 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 | ||
qed | qed | ||
− | (*pablucoto*) | + | (*pablucoto pabrodmac marcarmor13 jeamacpov *) |
− | |||
lemma "todos (λy. y=x) (copia n x)" | lemma "todos (λy. y=x) (copia n x)" | ||
proof (induct n) | proof (induct n) | ||
Línea 276: | Línea 302: | ||
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 (Suc 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 (Suc 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 | ||
qed | qed | ||
− | (* anaprarod *) | + | (* anaprarod marpoldia1 *) |
lemma "todos (λy. y=x) (copia n x)" | lemma "todos (λy. y=x) (copia n x)" | ||
proof (induct n) | proof (induct n) | ||
show "todos (λy. y = x) (copia 0 x)" by simp | show "todos (λy. y = x) (copia 0 x)" by simp | ||
− | + | next | |
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 | ||
finally show "todos (λy. y = x) (copia (Suc n) x)" using HI by simp | finally show "todos (λy. y = x) (copia (Suc n) x)" using HI by simp | ||
+ | qed | ||
+ | |||
+ | (* paupeddeg antsancab1 *) | ||
+ | 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]) ∧ | ||
+ | (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 | ||
+ | |||
+ | (* dancorgar *) | ||
+ | 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) (x#(copia n x))" using HI by simp | ||
+ | also have "... = (copia (Suc n) x)" by simp | ||
+ | finally show "todos (λy. y=x) (copia (Suc n) x)" by simp | ||
+ | qed | ||
+ | |||
+ | (* fracorjim1 *) | ||
+ | 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)) | ||
+ | finally 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 "... = ((λy. y=x) x ∧ todos (λy. y=x) (copia n x))" | ||
+ | by (simp only: todos.simps(2)) | ||
+ | also have "... = True ∧ True" using HI by simp | ||
+ | finally show "todos (λy. y=x) (copia (Suc n) x)" by simp | ||
qed | qed | ||
Línea 329: | Línea 404: | ||
"factI n = factI' n 1" | "factI n = factI' n 1" | ||
− | (* fraortmoy | + | (* fraortmoy danrodcha serrodcal pablucoto wilmorort anaprarod |
+ | marpoldia1 juacabsou josgarsan rubgonmar paupeddeg pabrodmac bowma | ||
+ | fracorjim1 dancorgar antsancab1 *) | ||
lemma fact: "factI' n x = x * factR n" | lemma fact: "factI' n x = x * factR n" | ||
proof (induct n) | proof (induct n) | ||
Línea 335: | Línea 412: | ||
next | next | ||
fix n | fix n | ||
− | assume H1 : | + | assume H1 : "factI' n x = x * factR n" |
− | have "factI' (Suc n) x = factI' n (x * Suc n)" by (simp only:factI'.simps(2)) | + | have "factI' (Suc n) x = factI' n (x * Suc n)" |
− | also have "... = (x * Suc n) * factR n" using H1 by simp (* no entiendo por qué no hace esto bien y luego todo funciona *) | + | 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 *) | ||
+ | (* Creo que es porque no le estás dando la hipótesis de inducción | ||
+ | sino otra que se le parece *) | ||
also have "... = x * factR (Suc n)" by (simp del: mult_Suc) | also have "... = x * factR (Suc n)" by (simp del: mult_Suc) | ||
finally show "factI' (Suc n) x = x * factR (Suc n)" by simp | finally show "factI' (Suc n) x = x * factR (Suc n)" by simp | ||
qed | qed | ||
− | (* crigomgom ivamenjim ferrenseg *) | + | (* crigomgom ivamenjim ferrenseg marcarmor13 jeamacpov anaprarod *) |
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 361: | Línea 442: | ||
(* migtermor *) | (* migtermor *) | ||
− | |||
lemma fact: "factI' n x = x * factR n" | lemma fact: "factI' n x = x * factR n" | ||
proof (induct n) | proof (induct n) | ||
Línea 371: | Línea 451: | ||
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 only: factI'.simps(2)) | + | have "factI' (Suc n) x = factI' n (x * Suc n)" |
+ | by (simp only: factI'.simps(2)) | ||
also have "factI' (Suc n) x = factI' n (x * Suc n)" by simp | also have "factI' (Suc n) x = factI' n (x * Suc n)" by simp | ||
have "factI' (Suc n) x = x * Suc n * factR n" using HI by simp | have "factI' (Suc n) x = x * Suc n * factR n" using HI by simp | ||
also have "... = x * factR (Suc n)" by (simp del: mult_Suc) | also have "... = x * factR (Suc n)" by (simp del: mult_Suc) | ||
− | finally show "factI' (Suc n) x = x * factR (Suc n)" by simp (* No entiendo por qué no acepta esto como demostrado *) | + | finally show "factI' (Suc n) x = x * factR (Suc n)" by simp |
+ | (* No entiendo por qué no acepta esto como demostrado *) | ||
+ | oops | ||
+ | |||
+ | (* ivamenjim *) | ||
+ | |||
+ | 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 x | ||
+ | assume HI: "⋀x. factI' n x = x * factR n" | ||
+ | have "factI' (Suc n) x = factI' n (Suc n)*x" by simp | ||
+ | (* Primero (Suc n) y luego multiplicado por x, como dice el enunciado del ejercicio *) | ||
+ | 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 | qed | ||
Línea 384: | Línea 482: | ||
------------------------------------------------------------------- *} | ------------------------------------------------------------------- *} | ||
− | (* fraortmoy danrodcha crigomgom ivamenjim serrodcal ferrenseg pablucoto wilmorort rubgonmar *) | + | (* fraortmoy danrodcha crigomgom ivamenjim serrodcal ferrenseg pablucoto |
+ | wilmorort rubgonmar marpoldia1 juacabsou josgarsan paupeddeg bowma anaprarod marcarmor13 antsancab1 *) | ||
corollary "factI n = factR n" | corollary "factI n = factR n" | ||
proof - | proof - | ||
Línea 396: | Línea 495: | ||
proof - | proof - | ||
have "factI n = factI' n 1" by simp | have "factI n = factI' n 1" by simp | ||
− | also have "... = 1 * factI' n 1" using fact by simp | + | also have "... = 1 * factI' n 1" by simp |
− | finally show "factI n = factR n" by simp | + | also have "... = 1 * factR n" using fact by simp |
+ | finally show "factI n = factR n" by simp | ||
+ | qed | ||
+ | |||
+ | (* pabrodmac *) | ||
+ | corollary "factI n = factR n" | ||
+ | proof - | ||
+ | have "factI n = factI' n 1" by simp | ||
+ | also have "... = 1 * factR n" using fact by simp | ||
+ | finally show "factI n = factR n" by simp | ||
qed | qed | ||
+ | (* dancorgar *) | ||
+ | corollary "factI n = factR n" | ||
+ | proof (induct n) | ||
+ | show "factI 0 = factR 0" by simp | ||
+ | next | ||
+ | fix n | ||
+ | assume "factI n = factR n" | ||
+ | have "factI (Suc n) = factI' (Suc n) 1" by simp | ||
+ | also have "... = 1 * factR (Suc n)" by (simp add: fact) | ||
+ | also have "... = factR (Suc n)" by simp | ||
+ | finally show "factI (Suc n) = factR (Suc n)" by simp | ||
+ | qed | ||
text {* --------------------------------------------------------------- | text {* --------------------------------------------------------------- | ||
Línea 418: | Línea 538: | ||
------------------------------------------------------------------- *} | ------------------------------------------------------------------- *} | ||
− | (* crigomgom fraortmoy rubgonmar ivamenjim serrodcal pablucoto wilmorort anaprarod*) | + | (* crigomgom fraortmoy rubgonmar ivamenjim serrodcal pablucoto wilmorort |
+ | anaprarod marpoldia1 juacabsou paupeddeg pabrodmac lucnovdos | ||
+ | dancorgar bowma marcarmor13 antsancab1 *) | ||
lemma "amplia xs y = xs @ [y]" | lemma "amplia xs y = xs @ [y]" | ||
proof (induct xs) | proof (induct xs) | ||
Línea 431: | Línea 553: | ||
qed | qed | ||
− | (*danrodcha ferrenseg *) | + | (* danrodcha ferrenseg *) |
lemma "amplia xs y = xs @ [y]" (is "?P xs") | lemma "amplia xs y = xs @ [y]" (is "?P xs") | ||
proof (induct xs) | proof (induct xs) | ||
Línea 452: | Línea 574: | ||
fix x xs | fix x xs | ||
assume HI: "amplia xs y = xs @ [y]" | assume HI: "amplia xs y = xs @ [y]" | ||
− | have "amplia (x # xs) y = x # amplia xs y" by (simp only: amplia.simps(2)) | + | 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])" using HI by simp | ||
also have "... = (x # xs) @ [y]" by simp | also have "... = (x # xs) @ [y]" by simp |
Revisión actual del 13:11 16 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"
(* crigomgom fraortmoy marpoldia1 ivamenjim serrodcal rubgonmar
ferrenseg juacabsou wilmorort josgarsan lucnovdos paupeddeg bowma
pabrodmac dancorgar antsancab1 *)
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
(* migtermor *)
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 only: sumaImpares.simps(2))
also have "sumaImpares (Suc n) = n*n + 2*n+1" using HI by simp
finally show "sumaImpares (Suc n) = (Suc n)*(Suc n)" by simp
qed
(* fracorjim1: Especifico la regla de simplificación y el paso del
desarrollo al cuadrado para hacerlo más legible. *)
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 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
(* palucoto anaprarod marcarmor13 jeamacpov *)
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)
------------------------------------------------------------------- *}
(*crigomgom ivamenjim danrodcha serrodcal rubgonmar ferrenseg juacabsou
wilmorort anaprarod marpoldia1 josgarsan paupeddeg pabrodmac bowma
dancorgar antsancab1 *)
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
(* migtermor *)
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 only: sumaPotenciasDeDosMasUno.simps(2))
also have "sumaPotenciasDeDosMasUno (Suc n) = 2^(n+1)+2^(n+1)"
using HI by simp
finally show "sumaPotenciasDeDosMasUno (Suc n) = 2^((Suc n)+1)"
by simp
qed
(* pablucoto marcarmor13 lucnovdos jeamacpov *)
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
finally show " sumaPotenciasDeDosMasUno (Suc n) = 2 ^ (Suc n + 1)"
by simp
qed
(* fracorjim1 - Hago explícita toda la manipulación algebraica. Quizás
excesivo.*)
lemma "sumaPotenciasDeDosMasUno n = 2^(n+1)"
proof (induction 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 only:sumaPotenciasDeDosMasUno.simps(2))
also have "... = 2 ^ (n + 1) + 2 ^ (n + 1)" using HI by simp
also have "... = 2 ^ (n + 1) * 2" by simp
also have "... = 2 ^ ((n + 1) + 1)" 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 ivamenjim serrodcal ferrenseg wilmorort juacabsou josgarsan
lucnovdos *)
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 rubgonmar *)
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
(* migtermor *)
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 "todos (λy. y=x) []" by (simp only: todos.simps(1))
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) (x#(copia n x)) =
((λy. y=x) x ∧ todos (λy. y=x) (copia n x))"
by (simp only: todos.simps(2))
also have "todos (λy. y=x) (copia (Suc n) x) = ((λy. y=x) x)"
using HI by simp
also have "((λy. y=x) x) = True" by simp
finally show "(todos (λy. y=x) (copia (Suc n) x))" by simp
qed
(*pablucoto pabrodmac marcarmor13 jeamacpov *)
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 (Suc n) x))"
by simp
also have "... = True" using HI by simp
finally show " todos (λy. y = x) (copia (Suc n) x) " by simp
qed
(* anaprarod marpoldia1 *)
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
finally show "todos (λy. y = x) (copia (Suc n) x)" using HI by simp
qed
(* paupeddeg antsancab1 *)
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]) ∧
(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
(* dancorgar *)
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) (x#(copia n x))" using HI by simp
also have "... = (copia (Suc n) x)" by simp
finally show "todos (λy. y=x) (copia (Suc n) x)" by simp
qed
(* fracorjim1 *)
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))
finally 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 "... = ((λy. y=x) x ∧ todos (λy. y=x) (copia n x))"
by (simp only: todos.simps(2))
also have "... = True ∧ 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"
(* fraortmoy danrodcha serrodcal pablucoto wilmorort anaprarod
marpoldia1 juacabsou josgarsan rubgonmar paupeddeg pabrodmac bowma
fracorjim1 dancorgar antsancab1 *)
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 *)
(* Creo que es porque no le estás dando la hipótesis de inducción
sino otra que se le parece *)
also have "... = x * factR (Suc n)" by (simp del: mult_Suc)
finally show "factI' (Suc n) x = x * factR (Suc n)" by simp
qed
(* crigomgom ivamenjim ferrenseg marcarmor13 jeamacpov anaprarod *)
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
(* migtermor *)
lemma fact: "factI' n x = x * factR n"
proof (induct n)
have "x * factR 0 = x" by (simp only: factR.simps(1))
also have "factI' 0 x = x" by (simp only: factI'.simps(1))
show "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 only: factI'.simps(2))
also have "factI' (Suc n) x = factI' n (x * Suc n)" by simp
have "factI' (Suc n) x = x * Suc n * factR n" using HI by simp
also have "... = x * factR (Suc n)" by (simp del: mult_Suc)
finally show "factI' (Suc n) x = x * factR (Suc n)" by simp
(* No entiendo por qué no acepta esto como demostrado *)
oops
(* ivamenjim *)
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 x
assume HI: "⋀x. factI' n x = x * factR n"
have "factI' (Suc n) x = factI' n (Suc n)*x" by simp
(* Primero (Suc n) y luego multiplicado por x, como dice el enunciado del ejercicio *)
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
text {* ---------------------------------------------------------------
Ejercicio 4.3. Escribir la demostración detallada de
factI n = factR n
------------------------------------------------------------------- *}
(* fraortmoy danrodcha crigomgom ivamenjim serrodcal ferrenseg pablucoto
wilmorort rubgonmar marpoldia1 juacabsou josgarsan paupeddeg bowma anaprarod marcarmor13 antsancab1 *)
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
(* migtermor *)
corollary "factI n = factR n"
proof -
have "factI n = factI' n 1" by simp
also have "... = 1 * factI' n 1" by simp
also have "... = 1 * factR n" using fact by simp
finally show "factI n = factR n" by simp
qed
(* pabrodmac *)
corollary "factI n = factR n"
proof -
have "factI n = factI' n 1" by simp
also have "... = 1 * factR n" using fact by simp
finally show "factI n = factR n" by simp
qed
(* dancorgar *)
corollary "factI n = factR n"
proof (induct n)
show "factI 0 = factR 0" by simp
next
fix n
assume "factI n = factR n"
have "factI (Suc n) = factI' (Suc n) 1" by simp
also have "... = 1 * factR (Suc n)" by (simp add: fact)
also have "... = factR (Suc n)" by simp
finally show "factI (Suc n) = factR (Suc 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 rubgonmar ivamenjim serrodcal pablucoto wilmorort
anaprarod marpoldia1 juacabsou paupeddeg pabrodmac lucnovdos
dancorgar bowma marcarmor13 antsancab1 *)
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 ferrenseg *)
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
(* migtermor *)
lemma "amplia xs y = xs @ [y]"
proof (induct xs)
have "amplia [] y = [y]" by (simp only: amplia.simps(1))
have "[] @ [y] = [y]" by simp
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
end