Acciones

Sol 11

De Lógica matemática y fundamentos (2018-19)

chapter {* R11: Razonamiento sobre programas en Isabelle/HOL *}
 
theory R11_sol
imports Main 
begin

text {* --------------------------------------------------------------- 
   En toda la relación de ejercicios las demostraciones han de realizarse
   de las formas siguientes:
    (*) detallada
    (*) estructurada
    (*) aplicativa
    (*) automática
  ------------------------------------------------------------------ *}
    
    
text {* --------------------------------------------------------------- 
  Ejercicio 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)"
 
value "sumaPotenciasDeDosMasUno 3"  = 16
 
text {* --------------------------------------------------------------- 
  Ejercicio 2. Demostrar que 
     sumaPotenciasDeDosMasUno n = 2^(n+1)
  ------------------------------------------------------------------- *}

lemma sumaPotenciasDeDosMasUno_d: 
  "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)"
  show "sumaPotenciasDeDosMasUno (Suc n) = 2 ^ (Suc n + 1)"
    proof -
      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 ?thesis .
    qed
  qed


lemma sumaPotenciasDeDosMasUno_e: 
  "sumaPotenciasDeDosMasUno n = 2^(n+1)"
oops

lemma sumaPotenciasDeDosMasUno_a: 
  "sumaPotenciasDeDosMasUno n = 2^(n+1)"
  apply (induct n)
   apply simp_all
  done


lemma sumaPotenciasDeDosMasUno_auto: 
  "sumaPotenciasDeDosMasUno n = 2^(n+1)"
 by (induct n) simp_all

 
text {* --------------------------------------------------------------- 
  Ejercicio 3. 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"
 
value "copia 3 x"  = [x,x,x]
 
text {* --------------------------------------------------------------- 
  Ejercicio 4. 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
  Nota: La conjunción se representa por ∧
  ----------------------------------------------------------------- *}
 
fun todos :: "('a ⇒ bool) ⇒ 'a list ⇒ bool" where
  "todos p []     = True"
| "todos p (x#xs) = (p x ∧ todos p xs)"
 
value "todos (λx. x>(1::nat)) [2,6,4]"  = True
value "todos (λx. x>(2::nat)) [2,6,4]"  = False
 
text {* --------------------------------------------------------------- 
  Ejercicio 5. Demostrar que todos los elementos de (copia n x) son
  iguales a x. 
  ------------------------------------------------------------------- *}
 
lemma todos_copia_d: "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
  

lemma todos_copia_e: "todos (λy. y=x) (copia n x)"
  oops
    
lemma todos_copia_a: "todos (λy. y=x) (copia n x)"
  apply (induct n)    
  apply simp_all
  done


lemma todos_copia_auto: "todos (λy. y=x) (copia n x)"
  by (induct n) simp_all

text {* --------------------------------------------------------------- 
  Ejercicio 6. 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"
 
value "factR 4"  = 24
 
text {* --------------------------------------------------------------- 
  Ejercicio 7. 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
  ------------------------------------------------------------------- *}
 
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"
   
  
 
lemma fact_d: "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 only: mult.assoc)
  also have "... = x * factR (Suc n)" by simp                   
  finally show "factI' (Suc n) x = x * factR (Suc n)" by simp
  qed
qed


lemma fact_e: "factI' n x = x * factR n"
oops

thm mult_Suc
thm mult.assoc  
  
lemma fact_a: "factI' n x = x * factR n"
  apply (induct n arbitrary: x)  
   apply simp
  apply (simp del: mult_Suc)
  done


lemma fact_auto: "factI' n x = x * factR n"
by (induct n arbitrary: x) (auto simp del: mult_Suc)
 
text {* --------------------------------------------------------------- 
  Ejercicio 8. Demostrar que
     factI n = factR n
  ------------------------------------------------------------------- *}
 
corollary fact_equiv: "factI n = factR n"
by (simp add: fact_a)
 
text {* --------------------------------------------------------------- 
  Ejercicio 9. 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"
 
value "amplia [d,a] t"  = [d,a,t]
 
text {* --------------------------------------------------------------- 
  Ejercicio 10. Demostrar que 
     amplia xs y = xs @ [y]
  ------------------------------------------------------------------- *}
 
lemma amplia_append_d: "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


lemma amplia_append_e: "amplia xs y = xs @ [y]"
  oops
    
lemma amplia_append_a: "amplia xs y = xs @ [y]"
   apply (induct xs)    
   apply simp_all
    done
 

lemma amplia_append_auto: "amplia xs y = xs @ [y]"
  by (induct xs) simp_all


text {* 
  --------------------------------------------------------------------- 
  Ejercicio 11. Definir la función 
     algunos :: ('a ⇒ bool) ⇒ 'a list ⇒ bool
  tal que (algunos p xs) se verifica si algunos elementos de la lista 
  xs cumplen la propiedad p. Por ejemplo, se verifica 
     algunos (λx. 1<length x) [[2,1,4],[3]]
     ¬algunos (λx. 1<length x) [[],[3]]"

  Nota: La función algunos es equivalente a la predefinida list_ex. 
  --------------------------------------------------------------------- 
*}

fun algunos  :: "('a ⇒ bool) ⇒ 'a list ⇒ bool" where
"algunos p [] = False"
| "algunos p (x#xs) = ((p x) ∨ (algunos p xs))"


  
text {*
  --------------------------------------------------------------------- 
  Ejercicio 12. Demostrar o refutar
     todos (λx. P x ∧ Q x) xs = (todos P xs ∧ todos Q xs)
  --------------------------------------------------------------------- 
*}
    
  
lemma conjCommuta4: "a ∧ b ∧ c ∧ d ⟷ a ∧ c ∧ b ∧ d"    
  apply (rule iffI)
   apply (rule conjI)
    apply (erule conjunct1)
   apply (rule conjI)
    apply (drule conjunct2)
    apply (drule conjunct2)
    apply (erule conjunct1) 
   apply (drule conjunct2)
   apply (frule conjunct1)
   apply (drule conjunct2)+
    apply (rule conjI, assumption+)
  apply auto
    done

lemma todos_conj_d: "todos (λx. P x ∧ Q x) xs = (todos P xs ∧ todos Q xs)"
  oops

lemma todos_conj_e: "todos (λx. P x ∧ Q x) xs = (todos P xs ∧ todos Q xs)"
  proof (induct xs)
  show "todos (λx. P x ∧ Q x) [] = (todos P [] ∧ todos Q [])" by simp
next
  fix a xs
  assume HI: "todos (λx. P x ∧ Q x) xs = (todos P xs ∧ todos Q xs)"
  thus "todos (λx. P x ∧ Q x) (a # xs) = (todos P (a # xs) ∧ todos Q (a # xs))" by auto
qed

thm mult_Suc
thm mult.assoc  
  

lemma todos_conj_a: "todos (λx. P x ∧ Q x) xs = (todos P xs ∧ todos Q xs)"
  apply (induct xs) 
  apply simp
  apply simp
  apply (simp add: conjCommuta4)
    done


lemma todos_conj_auto: "todos (λx. P x ∧ Q x) xs = (todos P xs ∧ todos Q xs)"
    by (induct xs) auto


text {*
  --------------------------------------------------------------------- 
  Ejercicio 13. Demostrar o refutar
     todos P (x @ y) = (todos P x ∧ todos P y)
  --------------------------------------------------------------------- 
*}

lemma todos_append_d: "todos P (xs @ ys) = (todos P xs ∧ todos P ys)"
proof (induct xs)
  show "todos P ([] @ ys) = (todos P [] ∧ todos P ys)" by simp
next
  fix a xs
  assume HI:"todos P (xs @ ys) = (todos P xs ∧ todos P ys)"
  show "todos P ((a # xs) @ ys) = (todos P (a # xs) ∧ todos P ys)" 
    proof -
      have "todos P ((a#xs) @ ys) = ((P a) ∧ todos P (xs@ys))" by simp
      also have "... = ((P a) ∧ todos P xs ∧ todos P ys)" using HI by simp
      also have "... = (todos P (a#xs) ∧ todos P ys)" by simp
      finally show ?thesis .
    qed
qed


lemma todos_append_e: "todos P (xs @ ys) = (todos P xs ∧ todos P ys)"
proof (induct xs)
  show "todos P ([] @ ys) = (todos P [] ∧ todos P ys)" by simp
next
  fix a xs
  assume "todos P (xs @ ys) = (todos P xs ∧ todos P ys)"
  thus "todos P ((a # xs) @ ys) = (todos P (a # xs) ∧ todos P ys)" by auto
qed


lemma todos_append_a: "todos P (xs @ ys) = (todos P xs ∧ todos P ys)"
apply (induct xs)
apply simp_all
done

lemma todos_append_auto: "todos P (x @ y) = (todos P x ∧ todos P y)"
by (induct x) simp_all

text {*
  --------------------------------------------------------------------- 
  Ejercicio 14. Demostrar o refutar
     todos P (rev xs) = todos P xs
  --------------------------------------------------------------------- 
*}

lemma todos_rev_d: "todos P (rev xs) = todos P xs"
proof (induct xs)
  show "todos P (rev []) = todos P []" by simp
next
  fix a xs
  assume HI:"todos P (rev xs) = todos P xs"
  show "todos P (rev (a # xs)) = todos P (a # xs)"
    proof -
      have "todos P (rev (a#xs)) = todos P ((rev xs)@[a])" by simp
      also have "... = (todos P (rev xs) ∧ todos P [a])"
        by (simp add: todos_append_a)
      also have "... = (todos P xs ∧ todos P [a])" using HI by simp
      also have "... = (todos P [a] ∧ todos P xs)" by auto
      also have "... = (P a ∧ todos P xs)" by simp
      also have "... = todos P (a#xs)" by simp
      finally show ?thesis .
    qed
qed


lemma todos_rev_e: "todos P (rev xs) = todos P xs"
proof (induct xs)
  show "todos P (rev []) = todos P []" by simp
next
  fix a xs
  assume "todos P (rev xs) = todos P xs"
  thus "todos P (rev (a#xs)) = todos P (a#xs)"
    by (auto simp add: todos_append_a)
qed


lemma todos_rev_a: "todos P (rev xs) = todos P xs"
oops

lemma todos_rev_auto: "todos P (rev xs) = todos P xs"
by (induct xs) (auto simp add: todos_append_a)

  
text {*
  --------------------------------------------------------------------- 
  Ejercicio 15. Demostrar o refutar:
    algunos (λx. P x ∧ Q x) xs = (algunos P xs ∧ algunos Q xs)
  --------------------------------------------------------------------- 
*}

lemma "algunos (λx. P x ∧ Q x) xs = (algunos P xs ∧ algunos Q xs)"
oops

text {*
  --------------------------------------------------------------------- 
  Ejercicio 16. Demostrar o refutar
     algunos P (map f xs) = algunos (P ∘ f) xs
  --------------------------------------------------------------------- 
*}

lemma algunos_map_d: "algunos P (map f xs) = algunos (P o f) xs"
proof (induct xs)
  show "algunos P (map f []) = algunos (P o f) []" by simp
next
  fix a xs
  assume HI: "algunos P (map f xs) = algunos (P o f) xs"
  show "algunos P (map f (a#xs)) = algunos (P o f) (a#xs)"
  proof -
    have "algunos P (map f (a#xs)) = algunos P ((f a)#(map f xs))" by simp
    also have "... = ((P (f a)) ∨ (algunos P (map f xs)))" by simp
    also have "... = (((P o f) a) ∨ (algunos (P o f) xs))" using HI by simp
    also have "... = algunos (P o f) (a#xs)" by simp
    finally show ?thesis .
  qed
qed


lemma algunos_map_e: "algunos P (map f xs) = algunos (P o f) xs"
proof (induct xs)
  show "algunos P (map f []) = algunos (P o f) []" by simp
next
  fix a xs
  assume "algunos P (map f xs) = algunos (P o f) xs"
  thus "algunos P (map f (a#xs)) = algunos (P o f) (a#xs)" by auto
qed

thm algunos.simps

lemma algunos_map_a: "algunos P (map f xs) = algunos (P o f) xs"
  apply (induct xs)
   apply simp_all
    done
  
thm conj_commute
    
      
      

lemma algunos_map_auto: "algunos P (map f xs) = algunos (P o f) xs"
by (induct xs) simp_all
  
text {*
  --------------------------------------------------------------------- 
  Ejercicio 17. Demostrar o refutar
     algunos P (xs @ ys) = (algunos P xs ∨ algunos P ys)
  --------------------------------------------------------------------- 
*}

lemma algunos_append_d: "algunos P (xs @ ys) = (algunos P xs ∨ algunos P ys)"
proof (induct xs)
  show "algunos P ([] @ ys) = (algunos P [] ∨ algunos P ys)" by simp
next
  fix a xs
  assume HI: "algunos P (xs @ ys) = (algunos P xs ∨ algunos P ys)"
  show "algunos P ((a#xs) @ ys) = (algunos P (a#xs) ∨ algunos P ys)"
  proof -
    have "algunos P ((a#xs) @ ys) = algunos P (a#(xs @ ys))" by simp
    also have "... = ((P a) ∨ algunos P (xs @ ys))" by simp
    also have "... = ((P a) ∨ algunos P xs ∨ algunos P ys)" using HI by simp
    also have "... = (algunos P (a#xs) ∨ algunos P ys)" by simp
    finally show ?thesis .
  qed
qed

thm Fun.comp_apply
  
lemma algunos_append_e: "algunos P (xs @ ys) = (algunos P xs ∨ algunos P ys)"
proof (induct xs)
  show "algunos P ([] @ ys) = (algunos P [] ∨ algunos P ys)" by simp
next
  fix a xs
  assume "algunos P (xs @ ys) = (algunos P xs ∨ algunos P ys)"
  thus "algunos P ((a#xs) @ ys) = (algunos P (a#xs) ∨ algunos P ys)"
    by auto
qed


lemma algunos_append_a: "algunos P (xs @ ys) = (algunos P xs ∨ algunos P ys)"
oops

lemma algunos_append_auto: "algunos P (xs @ ys) = (algunos P xs ∨ algunos P ys)"
by (induct xs) simp_all


text {*
  --------------------------------------------------------------------- 
  Ejercicio 18. Demostrar o refutar
     algunos P (rev xs) = algunos P xs
  --------------------------------------------------------------------- 
*}

lemma algunos_rev_d: "algunos P (rev xs) = algunos P xs"
proof (induct xs)
  show "algunos P (rev []) = algunos P []" by simp
next
  fix a xs
  assume HI: "algunos P (rev xs) = algunos P xs"
  show "algunos P (rev (a#xs)) = algunos P (a#xs)"
  proof -
    have "algunos P (rev (a#xs)) = algunos P ((rev xs) @ [a])" by simp
    also have "... = (algunos P (rev xs) ∨ algunos P [a])"
      by (simp add: algunos_append_auto)
    also have "... = (algunos P xs ∨ algunos P [a])" using HI by simp
    also have "... = (algunos P xs ∨ P a)" by simp
    also have "... = (P a ∨ algunos P xs)" by (rule disj_commute)
    also have "... = algunos P (a#xs)" by simp
    finally show ?thesis .
  qed
qed


lemma algunos_rev_e: "algunos P (rev xs) = algunos P xs"
proof (induct xs)
  show "algunos P (rev []) = algunos P []" by simp
next
  fix a xs
  assume "algunos P (rev xs) = algunos P xs"
  thus "algunos P (rev (a#xs)) = algunos P (a#xs)"
    by (auto simp add: algunos_append_auto)
qed


lemma algunos_rev_a: "algunos P (rev xs) = algunos P xs"
  apply (induct xs)
   apply simp
   apply (simp add: algunos_append_auto)
  apply (simp add: disj_commute)
   done 


lemma algunos_rev_auto: "algunos P (rev xs) = algunos P xs"
by (induct xs) (auto simp add: algunos_append_auto)


text {*
  --------------------------------------------------------------------- 
  Ejercicio 19. Encontrar un término no trivial Z tal que sea cierta la 
  siguiente ecuación:
     algunos (λx. P x ∨ Q x) xs = Z
  y demostrar la equivalencia de forma automática y detallada.
  --------------------------------------------------------------------- 
*}

lemma "algunos (λx. P x ∨ Q x) xs = (algunos P xs ∨ algunos Q xs)"
by (induct xs) auto

lemma "algunos (λx. P x ∨ Q x) xs = (algunos P xs ∨ algunos Q xs)"
  apply (induct xs)
   apply simp
  apply simp
    by auto
  

lemma "algunos (λx. P x ∨ Q x) xs = (algunos P xs ∨algunos Q xs)"
  proof (induct xs) 
    show "algunos (λx. P x ∨ Q x) [] = (algunos P []∨ algunos Q [])" by simp
   next
     fix a xs
     assume "algunos (λx. (P x ∨ Q x)) xs = (algunos P xs ∨ algunos Q xs)"
     thus "algunos (λx. P x ∨ Q x) (a#xs) = (algunos P(a#xs) ∨ algunos Q (a#xs))"
        by auto
    qed


text {*
  --------------------------------------------------------------------- 
  Ejercicio 20. Demostrar o refutar
     algunos P xs = (¬ todos (λx. (¬ P x)) xs)
  --------------------------------------------------------------------- 
*}


lemma algunos_no_todos_d: "algunos P xs = (¬ todos (λx. (¬ P x)) xs)"
proof (induct xs)
  show "algunos P [] = (¬ todos (λx. ¬ P x) [])" by simp
next
  fix a xs
  assume HI: "algunos P xs = (¬ todos (λx. ¬ P x) xs)"
  show "algunos P (a # xs) = (¬ todos (λx. ¬ P x) (a # xs))"
    proof -
      have "algunos P (a # xs) = ((P a) ∨ (algunos P xs))" by simp
      also have "... = ((P a) ∨ (¬ todos (λx. ¬ P x) xs))" using HI by simp
      also have "... = (¬ (¬ (P a) ∧ todos (λx. (¬ P x)) xs))" by simp
      also have "... = (¬ todos (λx. (¬ P x)) (a#xs))" by simp
      finally show ?thesis .
    qed
qed


lemma algunos_no_todos_e: "algunos P xs = (¬ todos (λx. (¬ P x)) xs)"
oops

lemma algunos_no_todos_a: "algunos P xs = (¬ todos (λx. (¬ P x)) xs)"
oops

lemma algunos_no_todos_auto: "algunos P xs = (¬ todos (λx. (¬ P x)) xs)"
by (induct xs) simp_all


 text {*
  --------------------------------------------------------------------- 
  Ejercicio 21. Definir la función
     estaEn :: 'a ⇒ 'a list ⇒ bool
  tal que (estaEn x xs) se verifica si el elemento x está en la lista
  xs. Por ejemplo, 
     estaEn (2::nat) [3,2,4] = True
     estaEn (1::nat) [3,2,4] = False
  --------------------------------------------------------------------- 
*}

fun estaEn :: "'a ⇒ 'a list ⇒ bool" where
"estaEn x [] = False"
| "estaEn x (a#xs) = (x=a ∨ estaEn x xs)"


text {*
  --------------------------------------------------------------------- 
  Ejercicio 22. Expresar la relación existente entre estaEn y algunos. 
  Demostrar dicha relación de forma automática y detallada.
  --------------------------------------------------------------------- 
*}

lemma estaEn_algunos_auto:
"estaEn y xs = algunos (λx. y=x) xs"
  by (induct xs) simp_all
   
lemma estaEn_algunos_a:
"estaEn y xs = algunos (λx. y=x) xs"    
  apply (induct xs)
   apply simp
  apply simp
    done


lemma estaEn_algunos_e:
"estaEn y xs = algunos (λ x. y=x) xs"
proof (induct xs)
  show "estaEn y [] = algunos (λ x. y=x) []" by simp
next
  fix a xs
  assume "estaEn y xs = algunos (λ x. y=x) xs"
  thus "estaEn y (a#xs) = algunos (λ x. y=x) (a#xs)" by simp
qed



text {* 
  --------------------------------------------------------------------- 
  Ejercicio 23. Definir la función
     sinDuplicados :: 'a list ⇒ bool
  tal que (sinDuplicados xs) se verifica si la lista xs no contiene
  duplicados. Por ejemplo,  
     sinDuplicados [1::nat,4,2]   = True
     sinDuplicados [1::nat,4,2,4] = False
  --------------------------------------------------------------------- 
*}

fun sinDuplicados :: "'a list ⇒ bool" where
"sinDuplicados [] = True"
| "sinDuplicados (a#xs) = ((¬ estaEn a xs) ∧ sinDuplicados xs)"


text {* 
  --------------------------------------------------------------------- 
  Ejercicio 24. Definir la función
     borraDuplicados :: 'a list ⇒ bool
  tal que (borraDuplicados xs) es la lista obtenida eliminando los
  elementos duplicados de la lista xs. Por ejemplo, 
     borraDuplicados [1::nat,2,4,2,3] = [1,4,2,3]

  Nota: La función borraDuplicados es equivalente a la predefinida 
  remdups. 
  --------------------------------------------------------------------- 
*}


fun borraDuplicados :: "'a list ⇒ 'a list" where
"borraDuplicados [] = []"
| "borraDuplicados (a#xs) = (if estaEn a xs
  then borraDuplicados xs
  else (a#borraDuplicados xs))"



text {*
  --------------------------------------------------------------------- 
  Ejercicio 25. Demostrar o refutar
     length (borraDuplicados xs) ≤ length xs
  --------------------------------------------------------------------- 
*}

lemma length_borraDuplicados_d:
"length (borraDuplicados xs) ≤ length xs"
proof (induct xs)
  show "length (borraDuplicados []) ≤ length []" by simp
next
  fix a xs
  assume HI: "length (borraDuplicados xs) ≤ length xs"
  show "length (borraDuplicados (a#xs)) ≤ length (a#xs)"
  proof (cases)
    assume "estaEn a xs"
    hence "length (borraDuplicados (a#xs)) = length (borraDuplicados xs)"
      by simp
    also have "... ≤ length xs" using HI by simp
    also have "... ≤ length (a#xs)" by simp
    finally show ?thesis .
  next
    assume "(¬ estaEn a xs)"
    hence "length (borraDuplicados (a#xs)) = length (a#(borraDuplicados xs))"
      by simp
    also have "... = 1 + length (borraDuplicados xs)" by simp
    also have "... ≤ 1 + length xs" using HI by simp
    also have "... = length (a#xs)" by simp
    finally show ?thesis .
  qed
qed


lemma length_borraDuplicados_e:
"length (borraDuplicados xs) ≤ length xs"
proof (induct xs)
  show "length (borraDuplicados [])  ≤ length []" by simp
next
  fix a xs
  assume HI: "length (borraDuplicados xs) ≤ length xs"
  thus "length (borraDuplicados (a#xs)) ≤ length (a#xs)"
  proof (cases)
    assume "estaEn a xs"
    thus "length (borraDuplicados (a#xs)) ≤ length (a#xs)"
      using HI by auto
  next
    assume "(¬ estaEn a xs)"
    thus "length (borraDuplicados (a#xs)) ≤ length (a#xs)"
      using HI by auto
  qed
qed


lemma length_borraDuplicados_a:
"length (borraDuplicados xs) ≤ length xs"
oops

lemma length_borraDuplicados_auto:
"length (borraDuplicados xs) ≤ length xs"
by (induct xs) simp_all


text {*
  --------------------------------------------------------------------- 
  Ejercicio 26. Demostrar o refutar
     estaEn a (borraDuplicados xs) = estaEn a xs
  --------------------------------------------------------------------- 
*}

lemma estaEn_borraDuplicados_d:
"estaEn a (borraDuplicados xs) = estaEn a xs"
proof (induct xs)
  show "estaEn a (borraDuplicados []) = estaEn a []" by simp
next
  fix b xs
  assume HI: "estaEn a (borraDuplicados xs) = estaEn a xs"
  show "estaEn a (borraDuplicados (b#xs)) = estaEn a (b#xs)"
  proof (rule iffI)
    assume c1: "estaEn a (borraDuplicados (b#xs))"
    show "estaEn a (b#xs)"
    proof (cases)
      assume "estaEn b xs"
      thus "estaEn a (b#xs)" using c1 HI by auto
    next
      assume "¬ estaEn b xs"
      thus "estaEn a (b#xs)" using c1 HI by auto
    qed
  next
    assume c2: "estaEn a (b#xs)"
    show "estaEn a (borraDuplicados (b#xs))"
    proof (cases)
      assume "a=b"
      thus "estaEn a (borraDuplicados (b#xs))" using HI by auto
    next
      assume "a ≠ b"
      thus "estaEn a (borraDuplicados (b#xs))" using `a ≠ b` c2 HI by auto
    qed
  qed
qed



lemma estaEn_borraDuplicados_e:
"estaEn a (borraDuplicados xs) = estaEn a xs"
proof (induct xs)
  show "estaEn a (borraDuplicados []) = estaEn a []" by simp
next
  fix b xs
  assume HI: "estaEn a (borraDuplicados xs) = estaEn a xs"
  show "estaEn a (borraDuplicados (b#xs)) = estaEn a (b#xs)"
  proof (rule iffI)
    assume c1: "estaEn a (borraDuplicados (b#xs))"
    show "estaEn a (b#xs)"
    proof (cases)
      assume "estaEn b xs"
      thus "estaEn a (b#xs)" using c1 HI by auto
    next
      assume "¬ estaEn b xs"
      thus "estaEn a (b#xs)" using c1 HI by auto
    qed
  next
    assume c2: "estaEn a (b#xs)"
    show "estaEn a (borraDuplicados (b#xs))"
    proof (cases)
      assume "a=b"
      thus "estaEn a (borraDuplicados (b#xs))" using HI by auto
    next
      assume "a ≠ b"
      thus "estaEn a (borraDuplicados (b#xs))" using `a ≠ b` c2 HI by auto
    qed
  qed
qed


lemma estaEn_borraDuplicados_a:
"estaEn a (borraDuplicados xs) = estaEn a xs"
oops

lemma estaEn_borraDuplicados_auto:
"estaEn a (borraDuplicados xs) = estaEn a xs"
by (induct xs) auto

text {*
  --------------------------------------------------------------------- 
  Ejercicio 27. Demostrar o refutar
     sinDuplicados (borraDuplicados xs)
  --------------------------------------------------------------------- 
*}

lemma sinDuplicados_borraDuplicados_d:
"sinDuplicados (borraDuplicados xs)"
proof (induct xs)
  show "sinDuplicados (borraDuplicados [])" by simp
next
  fix a xs
  assume HI: "sinDuplicados (borraDuplicados xs)"
  show "sinDuplicados (borraDuplicados (a#xs))"
  proof (cases)
    assume "estaEn a xs"
    thus "sinDuplicados (borraDuplicados (a#xs))" using HI by simp
  next
    assume "¬ estaEn a xs"
    hence "¬ (estaEn a xs) ∧ sinDuplicados (borraDuplicados xs)"
      using HI by simp
    hence "¬ estaEn a (borraDuplicados xs) ∧ sinDuplicados (borraDuplicados xs)"
      by (simp add: estaEn_borraDuplicados_auto)
    hence "sinDuplicados (a#borraDuplicados xs)" by simp
    thus "sinDuplicados (borraDuplicados (a#xs))"
      using `¬ estaEn a xs` by simp
  qed
qed


lemma sinDuplicados_borraDuplicados_e:
"sinDuplicados (borraDuplicados xs)"
proof (induct xs)
  show "sinDuplicados (borraDuplicados [])" by simp
next
  fix a xs
  assume HI: "sinDuplicados (borraDuplicados xs)"
  show "sinDuplicados (borraDuplicados (a#xs))"
  proof (cases)
    assume "estaEn a xs"
    thus "sinDuplicados (borraDuplicados (a#xs))" using HI by simp
  next
    assume "¬ estaEn a xs"
    thus "sinDuplicados (borraDuplicados (a#xs))"
      using `¬ estaEn a xs` HI by (auto simp add: estaEn_borraDuplicados_auto)
  qed
qed


lemma sinDuplicados_borraDuplicados_a:
"sinDuplicados (borraDuplicados xs)"
oops

lemma sinDuplicados_borraDuplicados_auto:
"sinDuplicados (borraDuplicados xs)"
by (induct xs) (auto simp add: estaEn_borraDuplicados_auto)


text {*
  --------------------------------------------------------------------- 
  Ejercicio 28. Demostrar o refutar:
    borraDuplicados (rev xs) = rev (borraDuplicados xs)
  --------------------------------------------------------------------- 
*}

lemma "borraDuplicados (rev xs) = rev (borraDuplicados xs)"
oops


end