Acciones

Diferencia entre revisiones de «Relación 6»

De Razonamiento automático (2013-14)

(Página creada con '<source lang="isar"> header {* R6: Sustitución, inversión y eliminación *} theory R6 imports Main begin section {* Sustitución, inversión y eliminación *} text {* -...')
 
Línea 17: Línea 17:
 
   ---------------------------------------------------------------------  
 
   ---------------------------------------------------------------------  
 
*}
 
*}
 +
 +
-- "maresccas4"
  
 
fun sust :: "'a ⇒ 'a ⇒ 'a list ⇒ 'a list" where
 
fun sust :: "'a ⇒ 'a ⇒ 'a list ⇒ 'a list" where
   "sust x y zs = undefined"
+
   "sust x y [] = []"
 +
| "sust x y (z#zs) = (if z=x then y#sust x y zs else z#sust x y zs)"
  
 
text {*
 
text {*
Línea 27: Línea 30:
 
   ---------------------------------------------------------------------   
 
   ---------------------------------------------------------------------   
 
*}
 
*}
 +
 +
-- "maresccas4"
  
 
lemma  
 
lemma  
 
   "sust x y (xs@ys) = (sust x y xs)@(sust x y ys)"
 
   "sust x y (xs@ys) = (sust x y xs)@(sust x y ys)"
oops
+
  by (induct xs) auto
  
 
text {*
 
text {*
Línea 40: Línea 45:
  
 
lemma sust_append:  
 
lemma sust_append:  
   "sust x y (xs@ys) = (sust x y xs)@(sust x y ys)"
+
   "sust x y (xs@ys) = (sust x y xs)@(sust x y ys)" (is "?P x y xs ys")
oops
+
proof (induct xs)
 +
show "?P x y [] ys" by simp
 +
next
 +
fix a :: "'a"
 +
fix xs :: "'a list"
 +
assume HI: "?P x y xs ys "
 +
show "?P x y (a#xs) ys"
 +
proof cases
 +
  assume C1: "a=x"
 +
  then have "sust x y ((a#xs)@ys) = y # sust x y (xs@ys)" by simp
 +
  also have "... = (y # sust x y xs) @ sust x y ys" using HI by simp
 +
  also have "... = sust x y (a#xs) @ sust x y ys" using C1 by simp
 +
  finally show "?P x y (a#xs) ys" by simp
 +
next
 +
  assume C2: "a≠x"
 +
  then have "sust x y ((a#xs)@ys) = a # sust x y (xs@ys)" by simp
 +
  also have "... = (a # sust x y xs) @ sust x y ys" using HI by simp
 +
  also have "... = sust x y (a#xs) @ sust x y ys" using C2 by simp
 +
  finally show "?P x y (a#xs) ys" by simp
 +
qed
 +
qed
  
 
text {*
 
text {*
Línea 49: Línea 74:
 
   ---------------------------------------------------------------------  
 
   ---------------------------------------------------------------------  
 
*}
 
*}
 +
 +
-- "maresccas4"
  
 
lemma  
 
lemma  
   "rev(sust x y zs) = sust x y (rev zs)"
+
   "rev(sust x y zs) = sust x y (rev zs)" (is "?P x y zs")
oops
+
by (induct zs) (auto simp add:sust_append)
  
 
text {*
 
text {*
Línea 60: Línea 87:
 
   ---------------------------------------------------------------------  
 
   ---------------------------------------------------------------------  
 
*}
 
*}
 +
 +
-- "maresccas4"
  
 
lemma rev_sust:  
 
lemma rev_sust:  
   "rev(sust x y zs) = sust x y (rev zs)"
+
   "rev(sust x y zs) = sust x y (rev zs)" (is "?P x y zs")
oops
+
proof (induct zs)
 +
show "?P x y []" by simp
 +
next
 +
fix z :: "'a"
 +
fix zs :: "'a list"
 +
assume HI: "?P x y zs"
 +
show "?P x y (z#zs)"
 +
proof cases
 +
  assume C1: "z=x"
 +
  then have "rev(sust x y (z#zs)) = rev(y # (sust x y zs))" by simp
 +
  also have "... = rev (sust x y zs) @ [y]" by (simp only: rev.simps(2))
 +
  also have "... = sust x y (rev zs) @ [y]" using HI by simp
 +
  also have "... = sust x y (rev (z#zs))" using C1 by (simp add:sust_append)
 +
  finally show "?P x y (z#zs)" by simp
 +
next
 +
  assume C2: "z≠x"
 +
  then have "rev(sust x y (z#zs)) = rev(z # (sust x y zs))" by simp
 +
  also have "... = rev (sust x y zs) @ [z]" by (simp only: rev.simps(2))
 +
  also have "... = sust x y (rev zs) @ [z]" using HI by simp
 +
  also have "... = sust x y (rev (z#zs))" using C2 by (simp add:sust_append)
 +
  finally show "?P x y (z#zs)" by simp
 +
qed
 +
qed
  
 
text {*
 
text {*
Línea 71: Línea 122:
 
   ---------------------------------------------------------------------  
 
   ---------------------------------------------------------------------  
 
*}
 
*}
 +
 +
-- "maresccas4"
  
 
lemma  
 
lemma  
 
   "sust x y (sust u v zs) = sust u v (sust x y zs)"
 
   "sust x y (sust u v zs) = sust u v (sust x y zs)"
 +
quickcheck
 
oops
 
oops
 +
(* Contraejemplo:
 +
  x = a⇣1
 +
  y = a⇣2
 +
  u = a⇣2
 +
  v = a⇣1
 +
  zs = [a⇣1]
 +
*)
  
 
text {*
 
text {*
Línea 82: Línea 143:
 
   ---------------------------------------------------------------------  
 
   ---------------------------------------------------------------------  
 
*}
 
*}
 +
 +
-- "maresccas4"
  
 
lemma  
 
lemma  
 
   "sust y z (sust x y zs) = sust x z zs"
 
   "sust y z (sust x y zs) = sust x z zs"
 +
  quickcheck
 
oops
 
oops
 +
(* Contraejemplo
 +
  y = a⇣1
 +
  z = a⇣2
 +
  x = a⇣2
 +
  zs = [a⇣1]
 +
*)
  
 
text {*
 
text {*
Línea 98: Línea 168:
 
   ---------------------------------------------------------------------  
 
   ---------------------------------------------------------------------  
 
*}
 
*}
 +
 +
-- "maresccas4"
  
 
fun borra :: "'a ⇒ 'a list ⇒ 'a list" where
 
fun borra :: "'a ⇒ 'a list ⇒ 'a list" where
   "borra x ys = undefined"
+
   "borra x [] = []"
 +
| "borra x (y#ys) = (if x=y then ys else y # borra x ys)"
  
 
text {*  
 
text {*  
Línea 111: Línea 184:
 
   ---------------------------------------------------------------------  
 
   ---------------------------------------------------------------------  
 
*}
 
*}
 +
 +
-- "maresccas4"
  
 
fun borraTodas :: "'a ⇒ 'a list ⇒ 'a list" where
 
fun borraTodas :: "'a ⇒ 'a list ⇒ 'a list" where
   "borraTodas x ys = undefined"
+
   "borraTodas x [] = []"
 +
| "borraTodas x (y#ys) = (if x=y then borraTodas x ys else y # borraTodas x ys)"
  
 
text {*
 
text {*
Línea 121: Línea 197:
 
   ---------------------------------------------------------------------  
 
   ---------------------------------------------------------------------  
 
*}
 
*}
 +
 +
-- "maresccas4"
  
 
lemma  
 
lemma  
 
   "borra x (borraTodas x xs) = borraTodas x xs"
 
   "borra x (borraTodas x xs) = borraTodas x xs"
oops
+
by (induct xs) auto
  
 
text {*
 
text {*
Línea 132: Línea 210:
 
   ---------------------------------------------------------------------  
 
   ---------------------------------------------------------------------  
 
*}
 
*}
 +
 +
-- "maresccas4"
  
 
lemma  
 
lemma  
   "borra x (borraTodas x xs) = borraTodas x xs"
+
   "borra x (borraTodas x xs) = borraTodas x xs" (is "?P x xs")
oops
+
proof (induct xs)
 +
show "?P x []" by simp
 +
next
 +
fix a :: "'a"
 +
fix xs :: "'a list"
 +
assume HI: "?P x xs"
 +
show "?P x (a#xs)"
 +
proof cases
 +
  assume C1: "x=a"
 +
  then have "borra x (borraTodas x (a#xs)) = borra x (borraTodas x xs)" by simp
 +
  also have "... = borraTodas x xs" using HI by simp
 +
  also have "... = borraTodas x (a#xs)" using C1 by simp
 +
  finally show "?P x (a#xs)" by simp
 +
next
 +
  assume C2: "x≠a"
 +
  then have "borra x (borraTodas x (a#xs)) = borra x (a # borraTodas x xs)" by simp
 +
  also have "... = a # borra x (borraTodas x xs)" using C2 by simp
 +
  also have "... = a # borraTodas x xs" using HI by simp
 +
  also have "... = borraTodas x (a#xs)" using C2 by simp
 +
  finally show "?P x (a#xs)" by simp
 +
qed
 +
qed
  
 
text {*
 
text {*
Línea 143: Línea 244:
 
   ---------------------------------------------------------------------  
 
   ---------------------------------------------------------------------  
 
*}
 
*}
 +
 +
-- "maresccas4"
  
 
lemma  
 
lemma  
 
   "borraTodas x (borraTodas x xs) = borraTodas x xs"
 
   "borraTodas x (borraTodas x xs) = borraTodas x xs"
oops
+
by (induct xs) auto
  
 
text {*
 
text {*
Línea 154: Línea 257:
 
   ---------------------------------------------------------------------  
 
   ---------------------------------------------------------------------  
 
*}
 
*}
 +
 +
-- "maresccas4"
  
 
lemma  
 
lemma  
   "borraTodas x (borraTodas x xs) = borraTodas x xs"
+
   "borraTodas x (borraTodas x xs) = borraTodas x xs" (is "?P x xs")
oops
+
proof (induct xs)
 +
show "?P x []" by simp
 +
next
 +
fix a :: "'a"
 +
fix xs :: "'a list"
 +
assume HI: "?P x xs"
 +
show "?P x (a#xs)"
 +
proof cases
 +
  assume C1: "x=a"
 +
  then have "borraTodas x (borraTodas x (a#xs)) = borraTodas x (borraTodas x xs)" by simp
 +
  also have "... = borraTodas x xs" using HI by simp
 +
  also have "... = borraTodas x (a#xs)" using C1 by simp
 +
  finally show "?P x (a#xs)" by simp
 +
next
 +
  assume C2: "x≠a"
 +
  then have "borraTodas x (borraTodas x (a#xs)) = a # borraTodas x (borraTodas x xs)" by simp
 +
  also have "... = a # borraTodas x xs" using HI by simp
 +
  also have "... = borraTodas x (a#xs)" using C2 by simp
 +
  finally show "?P x (a#xs)" by simp
 +
qed
 +
qed
  
 
text {*
 
text {*
Línea 165: Línea 290:
 
   ---------------------------------------------------------------------  
 
   ---------------------------------------------------------------------  
 
*}
 
*}
 +
 +
-- "maresccas4"
  
 
lemma  
 
lemma  
 
   "borraTodas x (borra x xs) = borraTodas x xs"
 
   "borraTodas x (borra x xs) = borraTodas x xs"
oops
+
by (induct xs) auto
 
      
 
      
 
text {*
 
text {*
Línea 176: Línea 303:
 
   ---------------------------------------------------------------------  
 
   ---------------------------------------------------------------------  
 
*}
 
*}
 +
 +
-- "maresccas4"
  
 
lemma  
 
lemma  
   "borraTodas x (borra x xs) = borraTodas x xs"
+
   "borraTodas x (borra x xs) = borraTodas x xs" (is "?P x xs")
oops
+
proof (induct xs)
 +
show "?P x []" by simp
 +
next
 +
fix a :: "'a"
 +
fix xs :: "'a list"
 +
assume HI: "?P x xs"
 +
show "?P x (a#xs)"
 +
proof cases
 +
  assume C1: "x=a"
 +
  then have "borraTodas x (borra x (a#xs)) = borraTodas x xs" by simp
 +
  also have "... = borraTodas x (a#xs)" using C1 by simp
 +
  finally show "?P x (a#xs)" by simp
 +
next
 +
  assume C2: "x≠a"
 +
  then have "borraTodas x (borra x (a#xs)) = a # borraTodas x (borra x xs)" by simp
 +
  also have "... = a # borraTodas x xs" using HI by simp
 +
  also have "... = borraTodas x (a#xs)" using C2 by simp
 +
  finally show "?P x (a#xs)" by simp
 +
qed
 +
qed
 
      
 
      
 
text {*
 
text {*

Revisión del 23:04 13 dic 2013

header {* R6: Sustitución, inversión y eliminación *}

theory R6
imports Main 
begin

section {* Sustitución, inversión y eliminación *}

text {* 
  --------------------------------------------------------------------- 
  Ejercicio 1. Definir la función 
     sust :: "'a ⇒ 'a ⇒ 'a list ⇒ 'a list"
  tal que (sust x y zs) es la lista obtenida sustituyendo cada
  occurrencia de x por y en la lista zs. Por ejemplo,
     sust (1::nat) 2 [1,2,3,4,1,2,3,4] = [2,2,3,4,2,2,3,4]
  --------------------------------------------------------------------- 
*}

-- "maresccas4"

fun sust :: "'a ⇒ 'a ⇒ 'a list ⇒ 'a list" where
  "sust x y [] = []"
| "sust x y (z#zs) = (if z=x then y#sust x y zs else z#sust x y zs)"

text {*
  --------------------------------------------------------------------- 
  Ejercicio 2.1. Demostrar o refutar automáticamente 
     sust x y (xs@ys) = (sust x y xs)@(sust x y ys)"
  ---------------------------------------------------------------------  
*}

-- "maresccas4"

lemma 
  "sust x y (xs@ys) = (sust x y xs)@(sust x y ys)"
  by (induct xs) auto

text {*
  --------------------------------------------------------------------- 
  Ejercicio 2.2. Demostrar o refutar detalladamente
     sust x y (xs@ys) = (sust x y xs)@(sust x y ys)"
  ---------------------------------------------------------------------  
*}

lemma sust_append: 
  "sust x y (xs@ys) = (sust x y xs)@(sust x y ys)" (is "?P x y xs ys")
proof (induct xs)
 show "?P x y [] ys" by simp
next 
 fix a :: "'a"
 fix xs :: "'a list"
 assume HI: "?P x y xs ys "
 show "?P x y (a#xs) ys"
 proof cases
  assume C1: "a=x"
  then have "sust x y ((a#xs)@ys) = y # sust x y (xs@ys)" by simp
  also have "... = (y # sust x y xs) @ sust x y ys" using HI by simp
  also have "... = sust x y (a#xs) @ sust x y ys" using C1 by simp
  finally show "?P x y (a#xs) ys" by simp
 next
  assume C2: "a≠x"
  then have "sust x y ((a#xs)@ys) = a # sust x y (xs@ys)" by simp
  also have "... = (a # sust x y xs) @ sust x y ys" using HI by simp
  also have "... = sust x y (a#xs) @ sust x y ys" using C2 by simp
  finally show "?P x y (a#xs) ys" by simp
 qed
qed

text {*
  --------------------------------------------------------------------- 
  Ejercicio 3.1. Demostrar o refutar automáticamente
     rev (sust x y zs) = sust x y (rev zs)
  --------------------------------------------------------------------- 
*}

-- "maresccas4"

lemma 
  "rev(sust x y zs) = sust x y (rev zs)" (is "?P x y zs")
by (induct zs) (auto simp add:sust_append)

text {*
  --------------------------------------------------------------------- 
  Ejercicio 3.2. Demostrar o refutar detalladamente
     rev (sust x y zs) = sust x y (rev zs)
  --------------------------------------------------------------------- 
*}

-- "maresccas4"

lemma rev_sust: 
  "rev(sust x y zs) = sust x y (rev zs)" (is "?P x y zs")
proof (induct zs)
 show "?P x y []" by simp
next
 fix z :: "'a"
 fix zs :: "'a list"
 assume HI: "?P x y zs"
 show "?P x y (z#zs)"
 proof cases
  assume C1: "z=x"
  then have "rev(sust x y (z#zs)) = rev(y # (sust x y zs))" by simp
  also have "... = rev (sust x y zs) @ [y]" by (simp only: rev.simps(2))
  also have "... = sust x y (rev zs) @ [y]" using HI by simp
  also have "... = sust x y (rev (z#zs))" using C1 by (simp add:sust_append)
  finally show "?P x y (z#zs)" by simp
 next
  assume C2: "z≠x"
  then have "rev(sust x y (z#zs)) = rev(z # (sust x y zs))" by simp
  also have "... = rev (sust x y zs) @ [z]" by (simp only: rev.simps(2))
  also have "... = sust x y (rev zs) @ [z]" using HI by simp
  also have "... = sust x y (rev (z#zs))" using C2 by (simp add:sust_append)
  finally show "?P x y (z#zs)" by simp
 qed
qed

text {*
  --------------------------------------------------------------------- 
  Ejercicio 4. Demostrar o refutar:
     sust x y (sust u v zs) = sust u v (sust x y zs)
  --------------------------------------------------------------------- 
*}

-- "maresccas4"

lemma 
  "sust x y (sust u v zs) = sust u v (sust x y zs)"
quickcheck
oops
(* Contraejemplo:
  x = a⇣1
  y = a⇣2
  u = a⇣2
  v = a⇣1
  zs = [a⇣1]
*)

text {*
  --------------------------------------------------------------------- 
  Ejercicio 5. Demostrar o refutar:
     sust y z (sust x y zs) = sust x z zs
  --------------------------------------------------------------------- 
*}

-- "maresccas4"

lemma 
  "sust y z (sust x y zs) = sust x z zs"
  quickcheck
oops
(* Contraejemplo
  y = a⇣1
  z = a⇣2
  x = a⇣2
  zs = [a⇣1]
*)

text {*
  --------------------------------------------------------------------- 
  Ejercicio 6. Definir la función
     borra :: "'a ⇒ 'a list ⇒ 'a list"
  tal que (borra x ys) es la lista obtenida borrando la primera
  ocurrencia del elemento x en la lista ys. Por ejemplo,
     borra (2::nat) [1,2,3,2] = [1,3,2]

  Nota: La función borra es equivalente a la predefinida remove1. 
  --------------------------------------------------------------------- 
*}

-- "maresccas4"

fun borra :: "'a ⇒ 'a list ⇒ 'a list" where
  "borra x [] = []"
| "borra x (y#ys) = (if x=y then ys else y # borra x ys)"

text {* 
  --------------------------------------------------------------------- 
  Ejercicio 7. Definir la función
     borraTodas :: "'a ⇒ 'a list ⇒ 'a list"
  tal que (borraTodas x ys) es la lista obtenida borrando todas las
  ocurrencias del elemento x en la lista ys. Por ejemplo,
     borraTodas (2::nat) [1,2,3,2] = [1,3]
  --------------------------------------------------------------------- 
*}

-- "maresccas4"

fun borraTodas :: "'a ⇒ 'a list ⇒ 'a list" where
  "borraTodas x [] = []"
| "borraTodas x (y#ys) = (if x=y then borraTodas x ys else y # borraTodas x ys)"

text {*
  --------------------------------------------------------------------- 
  Ejercicio 8.1. Demostrar o refutar automáticamente
     borra x (borraTodas x xs) = borraTodas x xs
  --------------------------------------------------------------------- 
*}

-- "maresccas4"

lemma 
  "borra x (borraTodas x xs) = borraTodas x xs"
by (induct xs) auto

text {*
  --------------------------------------------------------------------- 
  Ejercicio 8.2. Demostrar o refutar detalladamente 
     borra x (borraTodas x xs) = borraTodas x xs
  --------------------------------------------------------------------- 
*}

-- "maresccas4"

lemma 
  "borra x (borraTodas x xs) = borraTodas x xs" (is "?P x xs")
proof (induct xs)
 show "?P x []" by simp
next
 fix a :: "'a"
 fix xs :: "'a list"
 assume HI: "?P x xs"
 show "?P x (a#xs)"
 proof cases
  assume C1: "x=a"
  then have "borra x (borraTodas x (a#xs)) = borra x (borraTodas x xs)" by simp
  also have "... = borraTodas x xs" using HI by simp
  also have "... = borraTodas x (a#xs)" using C1 by simp
  finally show "?P x (a#xs)" by simp
 next
  assume C2: "x≠a"
  then have "borra x (borraTodas x (a#xs)) = borra x (a # borraTodas x xs)" by simp
  also have "... = a # borra x (borraTodas x xs)" using C2 by simp
  also have "... = a # borraTodas x xs" using HI by simp
  also have "... = borraTodas x (a#xs)" using C2 by simp
  finally show "?P x (a#xs)" by simp
 qed
qed

text {*
  --------------------------------------------------------------------- 
  Ejercicio 9.1. Demostrar o refutar automáticamente
     borraTodas x (borraTodas x xs) = borraTodas x xs
  --------------------------------------------------------------------- 
*}

-- "maresccas4"

lemma 
  "borraTodas x (borraTodas x xs) = borraTodas x xs"
by (induct xs) auto

text {*
  --------------------------------------------------------------------- 
  Ejercicio 9.2. Demostrar o refutar detalladamente
     borraTodas x (borraTodas x xs) = borraTodas x xs
  --------------------------------------------------------------------- 
*}

-- "maresccas4"

lemma 
  "borraTodas x (borraTodas x xs) = borraTodas x xs" (is "?P x xs")
proof (induct xs)
 show "?P x []" by simp
next
 fix a :: "'a"
 fix xs :: "'a list"
 assume HI: "?P x xs"
 show "?P x (a#xs)"
 proof cases
  assume C1: "x=a"
  then have "borraTodas x (borraTodas x (a#xs)) = borraTodas x (borraTodas x xs)" by simp
  also have "... = borraTodas x xs" using HI by simp
  also have "... = borraTodas x (a#xs)" using C1 by simp
  finally show "?P x (a#xs)" by simp
 next
  assume C2: "x≠a"
  then have "borraTodas x (borraTodas x (a#xs)) = a # borraTodas x (borraTodas x xs)" by simp
  also have "... = a # borraTodas x xs" using HI by simp
  also have "... = borraTodas x (a#xs)" using C2 by simp
  finally show "?P x (a#xs)" by simp
 qed
qed

text {*
  --------------------------------------------------------------------- 
  Ejercicio 10.1. Demostrar o refutar automáticamente
     borraTodas x (borra x xs) = borraTodas x xs
  --------------------------------------------------------------------- 
*}

-- "maresccas4"

lemma 
  "borraTodas x (borra x xs) = borraTodas x xs"
by (induct xs) auto
    
text {*
  --------------------------------------------------------------------- 
  Ejercicio 10.2. Demostrar o refutar detalladamente
     borraTodas x (borra x xs) = borraTodas x xs
  --------------------------------------------------------------------- 
*}

-- "maresccas4"

lemma 
  "borraTodas x (borra x xs) = borraTodas x xs" (is "?P x xs")
proof (induct xs)
 show "?P x []" by simp
next
 fix a :: "'a"
 fix xs :: "'a list"
 assume HI: "?P x xs"
 show "?P x (a#xs)"
 proof cases
  assume C1: "x=a"
  then have "borraTodas x (borra x (a#xs)) = borraTodas x xs" by simp
  also have "... = borraTodas x (a#xs)" using C1 by simp
  finally show "?P x (a#xs)" by simp
 next
  assume C2: "x≠a"
  then have "borraTodas x (borra x (a#xs)) = a # borraTodas x (borra x xs)" by simp
  also have "... = a # borraTodas x xs" using HI by simp
  also have "... = borraTodas x (a#xs)" using C2 by simp
  finally show "?P x (a#xs)" by simp
 qed
qed
    
text {*
  --------------------------------------------------------------------- 
  Ejercicio 11.1. Demostrar o automáticamente
     borra x (borra y xs) = borra y (borra x xs)
  --------------------------------------------------------------------- 
*}

lemma 
  "borra x (borra y xs) = borra y (borra x xs)"
oops

text {*
  --------------------------------------------------------------------- 
  Ejercicio 11.2. Demostrar o refutar detalladamente
     borra x (borra y xs) = borra y (borra x xs)
  --------------------------------------------------------------------- 
*}

lemma 
  "borra x (borra y xs) = borra y (borra x xs)"
oops

text {*
  --------------------------------------------------------------------- 
  Ejercicio 12.1. Demostrar o refutar automáticamente
     borraTodas x (borra y xs) = borra y (borraTodas x xs)
  --------------------------------------------------------------------- 
*}

lemma 
  "borraTodas x (borra y xs) = borra y (borraTodas x xs)"
oops

text {*
  --------------------------------------------------------------------- 
  Ejercicio 12.2. Demostrar o refutar detalladamente
     borraTodas x (borra y xs) = borra y (borraTodas x xs)
  --------------------------------------------------------------------- 
*}

lemma 
  "borraTodas x (borra y xs) = borra y (borraTodas x xs)"
oops

text {*
  --------------------------------------------------------------------- 
  Ejercicio 13. Demostrar o refutar:
     borra y (sust x y xs) = borra x xs
  --------------------------------------------------------------------- 
*}

lemma 
  "borra y (sust x y xs) = borra x xs"
oops

text {*
  --------------------------------------------------------------------- 
  Ejercicio 14. Demostrar o refutar:
     borraTodas y (sust x y xs) = borraTodas x xs"
  --------------------------------------------------------------------- 
*}

lemma 
  "borraTodas y (sust x y xs) = borraTodas x xs"
oops

text {*
  --------------------------------------------------------------------- 
  Ejercicio 15.1. Demostrar o refutar automáticamente
     sust x y (borraTodas x zs) = borraTodas x zs
  --------------------------------------------------------------------- 
*}

lemma 
  "sust x y (borraTodas x zs) = borraTodas x zs"
oops

text {*
  --------------------------------------------------------------------- 
  Ejercicio 15.2. Demostrar o refutar detalladamente
     sust x y (borraTodas x zs) = borraTodas x zs
  --------------------------------------------------------------------- 
*}

lemma 
  "sust x y (borraTodas x zs) = borraTodas x zs"
oops

text {*
  --------------------------------------------------------------------- 
  Ejercicio 16. Demostrar o refutar:
     sust x y (borraTodas z zs) = borraTodas z (sust x y zs)
  --------------------------------------------------------------------- 
*}

lemma 
  "sust x y (borraTodas z zs) = borraTodas z (sust x y zs)"
oops

text {*
  --------------------------------------------------------------------- 
  Ejercicio 17. Demostrar o refutar:
     rev (borra x xs) = borra x (rev xs)
  --------------------------------------------------------------------- 
*}

lemma 
  "rev (borra x xs) = borra x (rev xs)"
oops

text {*
  --------------------------------------------------------------------- 
  Ejercicio 18.1. Demostrar o refutar automáticamente
     borraTodas x (xs@ys) = (borraTodas x xs)@(borraTodas x ys)
  --------------------------------------------------------------------- 
*}

lemma 
  "borraTodas x (xs@ys) = (borraTodas x xs)@(borraTodas x ys)"
oops

text {*
  --------------------------------------------------------------------- 
  Ejercicio 18.2. Demostrar o refutar detalladamente
     borraTodas x (xs@ys) = (borraTodas x xs)@(borraTodas x ys)
  --------------------------------------------------------------------- 
*}

lemma 
  "borraTodas x (xs@ys) = (borraTodas x xs)@(borraTodas x ys)"
oops

text {*
  --------------------------------------------------------------------- 
  Ejercicio 19.1. Demostrar o refutar automáticamente
     rev (borraTodas x xs) = borraTodas x (rev xs)
  --------------------------------------------------------------------- 
*}

lemma 
  "rev (borraTodas x xs) = borraTodas x (rev xs)"
oops

text {*
  --------------------------------------------------------------------- 
  Ejercicio 19.2. Demostrar o refutar detalladamente
     rev (borraTodas x xs) = borraTodas x (rev xs)
  --------------------------------------------------------------------- 
*}

lemma 
  "rev (borraTodas x xs) = borraTodas x (rev xs)"
oops

end