Diferencia entre revisiones de «Relación 4»
De Razonamiento automático (2018-19)
(Página creada con «<source lang="isabelle"> chapter {* R4: Cuantificadores sobre listas *} theory R4_Cuantificadores_sobre_listas imports Main begin text {* ----------------------------…») |
|||
Línea 19: | Línea 19: | ||
*} | *} | ||
+ | (* pabalagon *) | ||
fun todos :: "('a ⇒ bool) ⇒ 'a list ⇒ bool" where | fun todos :: "('a ⇒ bool) ⇒ 'a list ⇒ bool" where | ||
− | "todos p xs = | + | "todos p [] = True" | |
+ | "todos p (x#xs) = (p x ∧ todos p xs)" | ||
text {* | text {* | ||
Línea 35: | Línea 37: | ||
*} | *} | ||
+ | (* pabalagon *) | ||
fun algunos :: "('a ⇒ bool) ⇒ 'a list ⇒ bool" where | fun algunos :: "('a ⇒ bool) ⇒ 'a list ⇒ bool" where | ||
− | "algunos p xs = | + | "algunos p [] = False" | |
+ | "algunos p (x#xs) = (p x ∨ algunos p xs)" | ||
text {* | text {* | ||
Línea 45: | Línea 49: | ||
*} | *} | ||
+ | (* pabalagon *) | ||
lemma "todos (λx. P x ∧ Q x) xs = (todos P xs ∧ todos Q xs)" | lemma "todos (λx. P x ∧ Q x) xs = (todos P xs ∧ todos Q xs)" | ||
− | + | by (induct xs) auto | |
text {* | text {* | ||
Línea 55: | Línea 60: | ||
*} | *} | ||
− | lemma "todos (λx. P x ∧ Q x) xs = (todos P xs ∧ todos Q xs)" | + | (* pabalagon *) |
− | + | lemma "todos (λx. P x ∧ Q x) xs = (todos P xs ∧ todos Q xs)" (is "?P P Q xs") | |
+ | proof (induct xs) | ||
+ | fix P Q | ||
+ | show "?P P Q []" by simp | ||
+ | next | ||
+ | fix a xs | ||
+ | assume HI: "?P P 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 ∧ todos P xs) ∧ (Q a ∧ todos Q xs))" by (simp add: HOL.conj_comms) | ||
+ | also have "... = (todos P (a#xs) ∧ (Q a ∧ todos Q xs))" by simp | ||
+ | finally show "todos (λx. P x ∧ Q x) (a#xs) = (todos P (a#xs) ∧ todos Q (a#xs))" by simp | ||
+ | qed | ||
text {* | text {* | ||
Línea 65: | Línea 82: | ||
*} | *} | ||
+ | (* pabalagon *) | ||
lemma "todos P (x @ y) = (todos P x ∧ todos P y)" | lemma "todos P (x @ y) = (todos P x ∧ todos P y)" | ||
− | + | by (induct x) auto | |
text {* | text {* | ||
Línea 75: | Línea 93: | ||
*} | *} | ||
+ | (* pabalagon *) | ||
lemma todos_append: | lemma todos_append: | ||
"todos P (x @ y) = (todos P x ∧ todos P y)" | "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 {* | text {* | ||
Línea 86: | Línea 113: | ||
*} | *} | ||
+ | (* pabalagon *) | ||
lemma "todos P (rev xs) = todos P xs" | lemma "todos P (rev xs) = todos P xs" | ||
− | + | by (induct xs) (simp_all add: HOL.conj_comms todos_append) | |
text {* | text {* | ||
Línea 96: | Línea 124: | ||
*} | *} | ||
+ | (* pabalagon *) | ||
lemma "todos P (rev xs) = todos P xs" | 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 | ||
+ | also have "... = (todos P xs ∧ P a)" by simp | ||
+ | also have "... = (P a ∧ todos P xs)" by (simp add: HOL.conj_comms) | ||
+ | finally show "todos P (rev (a#xs)) = todos P (a#xs)" by simp | ||
+ | qed | ||
text {* | text {* | ||
Línea 106: | Línea 146: | ||
*} | *} | ||
+ | (* pabalagon *) | ||
lemma "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 | + | nitpick |
+ | oops | ||
+ | |||
+ | (* Contraejemplo *) | ||
+ | value "algunos (λx. even x ∧ odd x) [1, 2::nat] ≠ | ||
+ | algunos even [1, 2::nat] ∧ algunos odd [1, 2::nat]" | ||
text {* | text {* | ||
Línea 116: | Línea 162: | ||
*} | *} | ||
+ | (* pabalagon *) | ||
lemma "algunos P (map f xs) = algunos (P o f) xs" | lemma "algunos P (map f xs) = algunos (P o f) xs" | ||
− | + | by (induct xs) auto | |
text {* | text {* | ||
--------------------------------------------------------------------- | --------------------------------------------------------------------- | ||
− | Ejercicio 7.2. Demostrar o refutar | + | Ejercicio 7.2. Demostrar o refutar detalladamente |
algunos P (map f xs) = algunos (P ∘ f) xs | algunos P (map f xs) = algunos (P ∘ f) xs | ||
--------------------------------------------------------------------- | --------------------------------------------------------------------- | ||
*} | *} | ||
+ | (* pabalagon *) | ||
lemma "algunos P (map f xs) = algunos (P o f) xs" | lemma "algunos P (map f xs) = algunos (P o f) xs" | ||
− | + | proof (induct xs) | |
+ | show "algunos P (map f []) = algunos (P ∘ 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)) = (P (f a) ∨ algunos P (map f xs))" by simp | ||
+ | also have "... = (P (f a) ∨ algunos (P ∘ f) xs)" using HI by simp | ||
+ | also have "... = ((P ∘ f) a ∨ algunos (P ∘ f) xs)" by simp | ||
+ | finally show "algunos P (map f (a#xs)) = algunos (P ∘ f) (a#xs)" by simp | ||
+ | qed | ||
text {* | text {* | ||
Línea 136: | Línea 193: | ||
*} | *} | ||
+ | (* pabalagon *) | ||
lemma "algunos P (xs @ ys) = (algunos P xs ∨ algunos P ys)" | lemma "algunos P (xs @ ys) = (algunos P xs ∨ algunos P ys)" | ||
− | + | by (induct xs) auto | |
text {* | text {* | ||
Línea 146: | Línea 204: | ||
*} | *} | ||
+ | (* pabalagon *) | ||
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)" | ||
− | + | 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 157: | Línea 224: | ||
*} | *} | ||
+ | (* pabalagon *) | ||
lemma "algunos P (rev xs) = algunos P xs" | lemma "algunos P (rev xs) = algunos P xs" | ||
− | + | by (induct xs) (simp_all add: HOL.disj_comms algunos_append) | |
text {* | text {* | ||
Línea 167: | Línea 235: | ||
*} | *} | ||
+ | (* pabalagon *) | ||
lemma "algunos P (rev xs) = algunos P xs" | 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 ∨ algunos P [a])" using HI by simp | ||
+ | also have "... = (algunos P xs ∨ P a)" by simp | ||
+ | also have "... = (P a ∨ algunos P xs)" by (simp add: HOL.disj_comms) | ||
+ | finally show "algunos P (rev (a#xs)) = algunos P (a#xs)" by simp | ||
+ | qed | ||
text {* | text {* | ||
Línea 186: | Línea 266: | ||
*} | *} | ||
+ | (* pabalagon *) | ||
lemma "algunos P xs = (¬ todos (λx. (¬ P x)) xs)" | lemma "algunos P xs = (¬ todos (λx. (¬ P x)) xs)" | ||
− | + | by (induct xs) auto | |
− | + | ||
text {* | text {* | ||
--------------------------------------------------------------------- | --------------------------------------------------------------------- | ||
Línea 196: | Línea 277: | ||
*} | *} | ||
+ | (* pabalagon *) | ||
lemma "algunos P xs = (¬ todos (λx. (¬ P x)) xs)" | lemma "algunos P xs = (¬ todos (λx. (¬ P x)) xs)" | ||
− | + | proof (induct xs) | |
+ | have "algunos P [] = False" by simp | ||
+ | also have "... = (¬ todos (λx. (¬ P x)) [])" by simp | ||
+ | finally show "algunos P [] = (¬ todos (λx. (¬ P x)) [])" by simp | ||
+ | next | ||
+ | fix a xs | ||
+ | assume HI: "algunos P xs = (¬ todos (λx. (¬ P x)) xs)" | ||
+ | have "(¬ todos (λx. (¬ P x)) (a#xs)) = (¬ ((¬ P a) ∧ todos (λx. (¬ P x)) xs))" by simp | ||
+ | also have "... = (P a ∨ ¬ todos (λx. (¬ P x)) xs)" by simp | ||
+ | also have "... = (P a ∨ algunos P xs)" using HI by simp | ||
+ | finally show "algunos P (a#xs) = (¬ todos (λx. (¬ P x)) (a#xs))" by simp | ||
+ | qed | ||
text {* | text {* | ||
Línea 210: | Línea 303: | ||
*} | *} | ||
+ | (* pabalagon *) | ||
fun estaEn :: "'a ⇒ 'a list ⇒ bool" where | fun estaEn :: "'a ⇒ 'a list ⇒ bool" where | ||
− | "estaEn x xs = | + | "estaEn x [] = False" | |
+ | "estaEn x (y#xs) = (x=y ∨ estaEn x xs)" | ||
text {* | text {* | ||
Línea 219: | Línea 314: | ||
--------------------------------------------------------------------- | --------------------------------------------------------------------- | ||
*} | *} | ||
+ | |||
+ | (* pabalagon *) | ||
+ | lemma "estaEn x xs = algunos (λa. x=a) xs" | ||
+ | by (induct xs) auto | ||
+ | |||
+ | (* pabalagon *) | ||
+ | lemma "estaEn x xs = algunos (λa. x=a) xs" | ||
+ | proof (induct xs) | ||
+ | show "estaEn x [] = algunos (λa. x=a) []" by simp | ||
+ | next | ||
+ | fix y xs | ||
+ | assume HI: "estaEn x xs = algunos (λa. x=a) xs" | ||
+ | have "estaEn x (y#xs) = (x=y ∨ estaEn x xs)" by simp | ||
+ | also have "... = (x=y ∨ algunos (λa. x=a) xs)" using HI by simp | ||
+ | finally show "estaEn x (y#xs) = algunos (λa. x=a) (y#xs)" by simp | ||
+ | qed | ||
end | end | ||
</source> | </source> |
Revisión del 00:37 3 dic 2018
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.
---------------------------------------------------------------------
*}
(* pabalagon *)
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.
---------------------------------------------------------------------
*}
(* pabalagon *)
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)
---------------------------------------------------------------------
*}
(* pabalagon *)
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)
---------------------------------------------------------------------
*}
(* pabalagon *)
lemma "todos (λx. P x ∧ Q x) xs = (todos P xs ∧ todos Q xs)" (is "?P P Q xs")
proof (induct xs)
fix P Q
show "?P P Q []" by simp
next
fix a xs
assume HI: "?P P 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 ∧ todos P xs) ∧ (Q a ∧ todos Q xs))" by (simp add: HOL.conj_comms)
also have "... = (todos P (a#xs) ∧ (Q a ∧ todos Q xs))" by simp
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)
---------------------------------------------------------------------
*}
(* pabalagon *)
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)
---------------------------------------------------------------------
*}
(* pabalagon *)
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
---------------------------------------------------------------------
*}
(* pabalagon *)
lemma "todos P (rev xs) = todos P xs"
by (induct xs) (simp_all add: HOL.conj_comms todos_append)
text {*
---------------------------------------------------------------------
Ejercicio 5.2. Demostrar o refutar detalladamente
todos P (rev xs) = todos P xs
---------------------------------------------------------------------
*}
(* pabalagon *)
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
also have "... = (todos P xs ∧ P a)" by simp
also have "... = (P a ∧ todos P xs)" by (simp add: HOL.conj_comms)
finally show "todos P (rev (a#xs)) = todos P (a#xs)" by simp
qed
text {*
---------------------------------------------------------------------
Ejercicio 6. Demostrar o refutar:
algunos (λx. P x ∧ Q x) xs = (algunos P xs ∧ algunos Q xs)
---------------------------------------------------------------------
*}
(* pabalagon *)
lemma "algunos (λx. P x ∧ Q x) xs = (algunos P xs ∧ algunos Q xs)"
nitpick
oops
(* Contraejemplo *)
value "algunos (λx. even x ∧ odd x) [1, 2::nat] ≠
algunos even [1, 2::nat] ∧ algunos odd [1, 2::nat]"
text {*
---------------------------------------------------------------------
Ejercicio 7.1. Demostrar o refutar automáticamente
algunos P (map f xs) = algunos (P ∘ f) xs
---------------------------------------------------------------------
*}
(* pabalagon *)
lemma "algunos P (map f xs) = algunos (P o f) xs"
by (induct xs) auto
text {*
---------------------------------------------------------------------
Ejercicio 7.2. Demostrar o refutar detalladamente
algunos P (map f xs) = algunos (P ∘ f) xs
---------------------------------------------------------------------
*}
(* pabalagon *)
lemma "algunos P (map f xs) = algunos (P o f) xs"
proof (induct xs)
show "algunos P (map f []) = algunos (P ∘ 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)) = (P (f a) ∨ algunos P (map f xs))" by simp
also have "... = (P (f a) ∨ algunos (P ∘ f) xs)" using HI by simp
also have "... = ((P ∘ f) a ∨ algunos (P ∘ f) xs)" 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)
---------------------------------------------------------------------
*}
(* pabalagon *)
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)
---------------------------------------------------------------------
*}
(* pabalagon *)
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
---------------------------------------------------------------------
*}
(* pabalagon *)
lemma "algunos P (rev xs) = algunos P xs"
by (induct xs) (simp_all add: HOL.disj_comms algunos_append)
text {*
---------------------------------------------------------------------
Ejercicio 9.2. Demostrar o refutar detalladamente
algunos P (rev xs) = algunos P xs
---------------------------------------------------------------------
*}
(* pabalagon *)
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 ∨ algunos P [a])" using HI by simp
also have "... = (algunos P xs ∨ P a)" by simp
also have "... = (P a ∨ algunos P xs)" by (simp add: HOL.disj_comms)
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.
---------------------------------------------------------------------
*}
text {*
---------------------------------------------------------------------
Ejercicio 11.1. Demostrar o refutar automáticamente
algunos P xs = (¬ todos (λx. (¬ P x)) xs)
---------------------------------------------------------------------
*}
(* pabalagon *)
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)
---------------------------------------------------------------------
*}
(* pabalagon *)
lemma "algunos P xs = (¬ todos (λx. (¬ P x)) xs)"
proof (induct xs)
have "algunos P [] = False" by simp
also have "... = (¬ todos (λx. (¬ P x)) [])" by simp
finally show "algunos P [] = (¬ todos (λx. (¬ P x)) [])" by simp
next
fix a xs
assume HI: "algunos P xs = (¬ todos (λx. (¬ P x)) xs)"
have "(¬ todos (λx. (¬ P x)) (a#xs)) = (¬ ((¬ P a) ∧ todos (λx. (¬ P x)) xs))" by simp
also have "... = (P a ∨ ¬ todos (λx. (¬ P x)) xs)" by simp
also have "... = (P a ∨ algunos P xs)" using HI 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
---------------------------------------------------------------------
*}
(* pabalagon *)
fun estaEn :: "'a ⇒ 'a list ⇒ bool" where
"estaEn x [] = False" |
"estaEn x (y#xs) = (x=y ∨ 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.
---------------------------------------------------------------------
*}
(* pabalagon *)
lemma "estaEn x xs = algunos (λa. x=a) xs"
by (induct xs) auto
(* pabalagon *)
lemma "estaEn x xs = algunos (λa. x=a) xs"
proof (induct xs)
show "estaEn x [] = algunos (λa. x=a) []" by simp
next
fix y xs
assume HI: "estaEn x xs = algunos (λa. x=a) xs"
have "estaEn x (y#xs) = (x=y ∨ estaEn x xs)" by simp
also have "... = (x=y ∨ algunos (λa. x=a) xs)" using HI by simp
finally show "estaEn x (y#xs) = algunos (λa. x=a) (y#xs)" by simp
qed
end