Acciones

Diferencia entre revisiones de «Relación 4»

De Razonamiento automático (2017-18)

Línea 210: Línea 210:
 
*}
 
*}
  
 +
(*anddonram*)
 
lemma algunos_append:
 
lemma algunos_append:
 
   "algunos P (xs @ ys) = (algunos P xs ∨ algunos P ys)"
 
   "algunos P (xs @ ys) = (algunos P xs ∨ algunos P ys)"
oops
+
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)"
 +
  have "algunos P ((a # xs) @ ys) = (P a ∨ algunos P (xs@ys) )" by simp
 +
  also have "... = (P a ∨ (algunos P xs ∨ algunos P ys))" using HI by simp
 +
  finally show "algunos P ((a # xs) @ ys) = (algunos P (a # xs) ∨ algunos P ys)" by simp
 +
qed
  
 
text {*
 
text {*
Línea 221: Línea 230:
 
*}
 
*}
  
 +
(*anddonram[conmutatividad or porque no sé hacerlo de otra forma] *)
 +
lemma or_comm: "(a ∨ b) = (b ∨ a)"
 +
  by (cases a) auto
 +
 +
(*anddonram*)
 
lemma "algunos P (rev xs) = algunos P xs"
 
lemma "algunos P (rev xs) = algunos P xs"
oops
+
  by (induct xs) (simp_all add: algunos_append or_comm)
 +
  
 
text {*
 
text {*
Línea 231: Línea 246:
 
*}
 
*}
  
 +
(*anddonram*)
 
lemma "algunos P (rev xs) = algunos P xs"
 
lemma "algunos P (rev xs) = algunos P xs"
oops
+
proof (induct xs)
 +
  show "algunos P (rev []) = algunos P []" by simp
 +
next
 +
  fix a xs
 +
  assume HI:" algunos P (rev xs) = algunos P xs"
 +
  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)
 +
  also have "... = (algunos P xs ∨ P a)" using HI by simp
 +
  also have "... = (P a ∨ algunos P xs)" by (simp add:or_comm)
 +
  finally show " algunos P (rev (a # xs)) = algunos P (a # xs) " by simp
 +
qed
  
 
text {*
 
text {*
Línea 242: Línea 268:
 
   ---------------------------------------------------------------------  
 
   ---------------------------------------------------------------------  
 
*}
 
*}
 +
 +
(*anddonram*)
 +
lemma "algunos (λx. P x ∨ Q x) xs = (algunos P xs ∨ algunos Q xs)"
 +
by (induct xs) auto
 +
 +
(*anddonram*)
 +
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 HI: " algunos (λx. P x ∨ Q x) xs = (algunos P xs ∨ algunos Q xs)"
 +
  have "algunos (λx. P x ∨ Q x) (a # xs) = (P a ∨ Q a ∨ algunos (λx. P x ∨ Q x) xs)"
 +
    by simp
 +
  also have "... = (P a ∨ Q a ∨ algunos P xs ∨ algunos Q xs)" using HI by simp
 +
  also have "... = (P a ∨ Q a ∨ algunos Q xs ∨ algunos P xs)" by (simp add: or_comm)
 +
  also have "... = (P a ∨  algunos Q (a#xs) ∨ algunos P xs)"  by simp
 +
  also have "... = (P a ∨  algunos P xs ∨ algunos Q (a#xs))" by (simp add: or_comm)
 +
  finally show "algunos (λx. P x ∨ Q x) (a # xs) = (algunos P (a # xs) ∨ algunos Q (a # xs))"
 +
    by simp
 +
qed
 +
  
 
text {*
 
text {*
Línea 250: Línea 298:
 
*}
 
*}
  
 +
(*anddonram*)
 
lemma "algunos P xs = (¬ todos (λx. (¬ P x)) xs)"
 
lemma "algunos P xs = (¬ todos (λx. (¬ P x)) xs)"
oops
+
by (induct xs) auto
   
+
 
text {*
 
text {*
 
   ---------------------------------------------------------------------  
 
   ---------------------------------------------------------------------  
Línea 260: Línea 309:
 
*}
 
*}
  
 +
(*anddonram*)
 
lemma "algunos P xs = (¬ todos (λx. (¬ P x)) xs)"
 
lemma "algunos P xs = (¬ todos (λx. (¬ P x)) xs)"
oops
+
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)"
 +
  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
 +
  finally show "  algunos P (a # xs) = (¬ todos (λx. ¬ P x) (a # xs))" by simp
 +
qed
 +
 
text {*
 
text {*
 
   ---------------------------------------------------------------------  
 
   ---------------------------------------------------------------------  
Línea 274: Línea 333:
 
*}
 
*}
  
 +
(*anddonram*)
 
fun estaEn :: "'a ⇒ 'a list ⇒ bool" where
 
fun estaEn :: "'a ⇒ 'a list ⇒ bool" where
   "estaEn x xs = undefined"
+
   "estaEn x [] = False"
 +
| "estaEn x (a#xs) = ((a=x) ∨ estaEn x xs)"
  
 
text {*
 
text {*
Línea 283: Línea 344:
 
   ---------------------------------------------------------------------  
 
   ---------------------------------------------------------------------  
 
*}
 
*}
 +
 +
(*anddonram*)
 +
lemma "estaEn x xs=algunos (λy.(y=x)) xs"
 +
  by (induct xs) auto
 +
 +
(*anddonram*)
 +
lemma "estaEn x xs=algunos (λy.(y=x)) xs"
 +
proof (induct xs)
 +
  show "estaEn x []=algunos (λy.(y=x)) []" by simp
 +
next
 +
  fix a xs
 +
  assume HI:"estaEn x xs = algunos (λy. y = x) xs"
 +
  have "estaEn x (a # xs) =((a=x) ∨ estaEn x xs)" by simp
 +
  also have "... = ((a=x) ∨ algunos (λy. y = x) xs)" using HI by simp
 +
  finally show "estaEn x (a # xs) = algunos (λy. y = x) (a # xs)" by simp
 +
qed
 +
  
 
end
 
end
 
</source>
 
</source>

Revisión del 13:20 2 dic 2017

chapter {* R4: Cuantificadores sobre listas *}

theory R4_Cuantificadores_sobre_listas
imports Main 
begin

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

  Nota: La función todos es equivalente a la predefinida list_all. 
  --------------------------------------------------------------------- 
*}

(*anddonram*)
fun todos :: "('a ⇒ bool) ⇒ 'a list ⇒ bool" where
  "todos p [] = True"
| "todos p (x#xs) = (p x ∧ todos p xs)"

text {* 
  --------------------------------------------------------------------- 
  Ejercicio 2. 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. 
  --------------------------------------------------------------------- 
*}

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

text {*
  --------------------------------------------------------------------- 
  Ejercicio 3.1. Demostrar o refutar automáticamente 
     todos (λx. P x ∧ Q x) xs = (todos P xs ∧ todos Q xs)
  --------------------------------------------------------------------- 
*}
 (*anddonram*)
lemma "todos (λx. P x ∧ Q x) xs = (todos P xs ∧ todos Q xs)"
  by (induct xs) auto
 

text {*
  --------------------------------------------------------------------- 
  Ejercicio 3.2. Demostrar o refutar detalladamente
     todos (λx. P x ∧ Q x) xs = (todos P xs ∧ todos Q xs)
  --------------------------------------------------------------------- 
*}


(*anddonram[conmutatividad and porque no sé hacerlo de otra forma] *)
lemma and_comm: "(a ∧ b) = (b ∧ a)"
  by (cases a) auto

(*anddonram*)
lemma "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)"
  have "todos (λx. P x ∧ Q x) (a # xs) = (P a ∧ Q a ∧ todos (λx. P x ∧ Q x) xs)" by simp
  also have "... = (P a ∧ Q a ∧ todos P xs ∧ todos Q xs)" using HI by simp
  also have "... = (P a ∧ Q a ∧ todos Q xs ∧ todos P xs)" by (simp add:and_comm)
  also have "... = (P a ∧ todos Q (a # xs) ∧ todos P xs)" by simp
  also have "... = (P a ∧ todos P xs ∧ Q a ∧ todos Q xs)" by (simp add:and_comm)
  finally show "todos (λx. P x ∧ Q x) (a # xs) = (todos P (a # xs) ∧ todos Q (a # xs))"
    by simp
qed

text {*
  --------------------------------------------------------------------- 
  Ejercicio 4.1. Demostrar o refutar automáticamente
     todos P (x @ y) = (todos P x ∧ todos P y)
  --------------------------------------------------------------------- 
*}

(*anddonram*)
lemma "todos P (x @ y) = (todos P x ∧ todos P y)"
  by (induct x) auto
 

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


(*anddonram*)
lemma todos_append:
  "todos P (x @ y) = (todos P x ∧ todos P y)"
proof (induct x)
  show "todos P ([] @ y) = (todos P [] ∧ todos P y)" by simp
next
  fix a x
  assume HI: "todos P (x @ y) = (todos P x ∧ todos P y)"
  have "todos P ((a # x) @ y) = (P a ∧ todos P (x @ y))" by simp
  also have "... = (P a ∧ todos P x ∧ todos P y)" using HI by simp
  finally show "todos P ((a # x) @ y) = (todos P (a # x) ∧ todos P y)" by simp
qed

text {*
  --------------------------------------------------------------------- 
  Ejercicio 5.1. Demostrar o refutar automáticamente 
     todos P (rev xs) = todos P xs
  --------------------------------------------------------------------- 
*}


(*anddonram*)
lemma "todos P (rev xs) = todos P xs"
  by (induct xs) (simp_all add: todos_append and_comm)

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


(*anddonram*)
lemma "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"
  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)
  also have "... =  (todos P xs ∧ todos P [a])" using HI by simp
  finally show "todos P (rev (a # xs)) = todos P (a # xs)" by (simp add:and_comm)
qed

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

(*anddonram Contraejemplo*)
 value "let xs=[True,False]    
  in (algunos (λx. (λx. (x=False)) x ∧ (λx. x) x) xs =
     (algunos (λx. (x=False)) xs ∧ algunos (λx. x) xs))"

text {*
  --------------------------------------------------------------------- 
  Ejercicio 7.1. Demostrar o refutar automáticamente 
     algunos P (map f xs) = algunos (P ∘ f) xs
  --------------------------------------------------------------------- 
*}

(*anddonram*)
lemma "algunos P (map f xs) = algunos (P o f) xs"
  by (induct xs) auto
 

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

 (*anddonram*)
lemma "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 ∘ f) xs"
  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 ∘ f) a ∨ algunos (P ∘ f) xs) " using HI by simp
  finally show "algunos P (map f (a # xs)) = algunos (P ∘ f) (a # xs)" by simp
qed

text {*
  --------------------------------------------------------------------- 
  Ejercicio 8.1. Demostrar o refutar automáticamente 
     algunos P (xs @ ys) = (algunos P xs ∨ algunos P ys)
  --------------------------------------------------------------------- 
*}

 (*anddonram*)
lemma "algunos P (xs @ ys) = (algunos P xs ∨ algunos P ys)"
  by (induct xs) auto
 

text {*
  --------------------------------------------------------------------- 
  Ejercicio 8.2. Demostrar o refutar detalladamente
     algunos P (xs @ ys) = (algunos P xs ∨ algunos P ys)
  --------------------------------------------------------------------- 
*}

 (*anddonram*)
lemma algunos_append:
  "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)"
  have "algunos P ((a # xs) @ ys) = (P a ∨ algunos P (xs@ys) )" by simp
  also have "... = (P a ∨ (algunos P xs ∨ algunos P ys))" using HI by simp
  finally show "algunos P ((a # xs) @ ys) = (algunos P (a # xs) ∨ algunos P ys)" by simp
qed

text {*
  --------------------------------------------------------------------- 
  Ejercicio 9.1. Demostrar o refutar automáticamente
     algunos P (rev xs) = algunos P xs
  --------------------------------------------------------------------- 
*}

(*anddonram[conmutatividad or porque no sé hacerlo de otra forma] *)
lemma or_comm: "(a ∨ b) = (b ∨ a)"
  by (cases a) auto

(*anddonram*)
lemma "algunos P (rev xs) = algunos P xs"
  by (induct xs) (simp_all add: algunos_append or_comm)
 

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

 (*anddonram*)
lemma "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"
  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)
  also have "... = (algunos P xs ∨ P a)" using HI by simp
  also have "... = (P a ∨ algunos P xs)" by (simp add:or_comm)
  finally show " algunos P (rev (a # xs)) = algunos P (a # xs) " by simp
qed

text {*
  --------------------------------------------------------------------- 
  Ejercicio 10. 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.
  --------------------------------------------------------------------- 
*}

 (*anddonram*)
lemma "algunos (λx. P x ∨ Q x) xs = (algunos P xs ∨ algunos Q xs)"
 by (induct xs) auto
 
(*anddonram*)
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 HI: " algunos (λx. P x ∨ Q x) xs = (algunos P xs ∨ algunos Q xs)"
  have "algunos (λx. P x ∨ Q x) (a # xs) = (P a ∨ Q a ∨ algunos (λx. P x ∨ Q x) xs)" 
    by simp
  also have "... = (P a ∨ Q a ∨ algunos P xs ∨ algunos Q xs)" using HI by simp
  also have "... = (P a ∨ Q a ∨ algunos Q xs ∨ algunos P xs)" by (simp add: or_comm)
  also have "... = (P a ∨  algunos Q (a#xs) ∨ algunos P xs)"  by simp 
  also have "... = (P a ∨  algunos P xs ∨ algunos Q (a#xs))" by (simp add: or_comm)
  finally show "algunos (λx. P x ∨ Q x) (a # xs) = (algunos P (a # xs) ∨ algunos Q (a # xs))"
    by simp
qed
 

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

 (*anddonram*)
lemma "algunos P xs = (¬ todos (λx. (¬ P x)) xs)"
 by (induct xs) auto
 
text {*
  --------------------------------------------------------------------- 
  Ejercicio 11.2. Demostrar o refutar datalladamente
     algunos P xs = (¬ todos (λx. (¬ P x)) xs)
  --------------------------------------------------------------------- 
*}

 (*anddonram*)
lemma "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)"
  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
  finally show "  algunos P (a # xs) = (¬ todos (λx. ¬ P x) (a # xs))" by simp
qed
 
text {*
  --------------------------------------------------------------------- 
  Ejercicio 12. Definir la funcion primitiva recursiva 
     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
  --------------------------------------------------------------------- 
*}

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

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

 (*anddonram*)
lemma "estaEn x xs=algunos (λy.(y=x)) xs"
  by (induct xs) auto

(*anddonram*)
lemma "estaEn x xs=algunos (λy.(y=x)) xs"
proof (induct xs)
  show "estaEn x []=algunos (λy.(y=x)) []" by simp
next
  fix a xs
  assume HI:"estaEn x xs = algunos (λy. y = x) xs"
  have "estaEn x (a # xs) =((a=x) ∨ estaEn x xs)" by simp
  also have "... = ((a=x) ∨ algunos (λy. y = x) xs)" using HI by simp
  finally show "estaEn x (a # xs) = algunos (λy. y = x) (a # xs)" by simp
qed


end