Cada familia de conjuntos P define una relación de forma que dos elementos están relacionados si algún conjunto de P contiene a ambos elementos. Se puede definir en Lean por
def relacion (P : set (set X)) (x y : X) :=
∃ A ∈ P, x ∈ A ∧ y ∈ A |
def relacion (P : set (set X)) (x y : X) :=
∃ A ∈ P, x ∈ A ∧ y ∈ A
Una familia de subconjuntos de X es una partición de X si cada elemento de X pertenece a un único conjunto de P y todos los elementos de P son no vacíos. Se puede definir en Lean por
def particion (P : set (set X)) : Prop :=
(∀ x, (∃ B ∈ P, x ∈ B ∧ ∀ C ∈ P, x ∈ C → B = C)) ∧ ∅ ∉ P |
def particion (P : set (set X)) : Prop :=
(∀ x, (∃ B ∈ P, x ∈ B ∧ ∀ C ∈ P, x ∈ C → B = C)) ∧ ∅ ∉ P
Demostrar que si P es una partición de X, entonces la relación definida por P es reflexiva.
Para ello, completar la siguiente teoría de Lean:
import tactic
variable {X : Type}
variable (P : set (set X))
def relacion (P : set (set X)) (x y : X) :=
∃ A ∈ P, x ∈ A ∧ y ∈ A
def particion (P : set (set X)) : Prop :=
(∀ x, (∃ B ∈ P, x ∈ B ∧ ∀ C ∈ P, x ∈ C → B = C)) ∧ ∅ ∉ P
example
(h : particion P)
: reflexive (relacion P) :=
sorry |
import tactic
variable {X : Type}
variable (P : set (set X))
def relacion (P : set (set X)) (x y : X) :=
∃ A ∈ P, x ∈ A ∧ y ∈ A
def particion (P : set (set X)) : Prop :=
(∀ x, (∃ B ∈ P, x ∈ B ∧ ∀ C ∈ P, x ∈ C → B = C)) ∧ ∅ ∉ P
example
(h : particion P)
: reflexive (relacion P) :=
sorry
[expand title=»Soluciones con Lean»]
import tactic
variable {X : Type}
variable (P : set (set X))
def relacion (P : set (set X)) (x y : X) :=
∃ A ∈ P, x ∈ A ∧ y ∈ A
def particion (P : set (set X)) : Prop :=
(∀ x, (∃ B ∈ P, x ∈ B ∧ ∀ C ∈ P, x ∈ C → B = C)) ∧ ∅ ∉ P
-- 1ª demostración
example
(h : particion P)
: reflexive (relacion P) :=
begin
unfold reflexive,
intro x,
unfold relacion,
unfold particion at h,
replace h : ∃ A ∈ P, x ∈ A ∧ ∀ B ∈ P, x ∈ B → A = B := h.1 x,
rcases h with ⟨A, hAP, hxA, -⟩,
use A,
repeat { split },
{ exact hAP, },
{ exact hxA, },
{ exact hxA, },
end
-- 2ª demostración
example
(h : particion P)
: reflexive (relacion P) :=
begin
intro x,
replace h : ∃ A ∈ P, x ∈ A ∧ ∀ B ∈ P, x ∈ B → A = B := h.1 x,
rcases h with ⟨A, hAP, hxA, -⟩,
use A,
repeat { split } ; assumption,
end
-- 3ª demostración
example
(h : particion P)
: reflexive (relacion P) :=
begin
intro x,
rcases (h.1 x) with ⟨A, hAP, hxA, -⟩,
use A,
repeat { split } ; assumption,
end
-- 4ª demostración
example
(h : particion P)
: reflexive (relacion P) :=
begin
intro x,
rcases (h.1 x) with ⟨A, hAP, hxA, -⟩,
use [A, ⟨hAP, hxA, hxA⟩],
end |
import tactic
variable {X : Type}
variable (P : set (set X))
def relacion (P : set (set X)) (x y : X) :=
∃ A ∈ P, x ∈ A ∧ y ∈ A
def particion (P : set (set X)) : Prop :=
(∀ x, (∃ B ∈ P, x ∈ B ∧ ∀ C ∈ P, x ∈ C → B = C)) ∧ ∅ ∉ P
-- 1ª demostración
example
(h : particion P)
: reflexive (relacion P) :=
begin
unfold reflexive,
intro x,
unfold relacion,
unfold particion at h,
replace h : ∃ A ∈ P, x ∈ A ∧ ∀ B ∈ P, x ∈ B → A = B := h.1 x,
rcases h with ⟨A, hAP, hxA, -⟩,
use A,
repeat { split },
{ exact hAP, },
{ exact hxA, },
{ exact hxA, },
end
-- 2ª demostración
example
(h : particion P)
: reflexive (relacion P) :=
begin
intro x,
replace h : ∃ A ∈ P, x ∈ A ∧ ∀ B ∈ P, x ∈ B → A = B := h.1 x,
rcases h with ⟨A, hAP, hxA, -⟩,
use A,
repeat { split } ; assumption,
end
-- 3ª demostración
example
(h : particion P)
: reflexive (relacion P) :=
begin
intro x,
rcases (h.1 x) with ⟨A, hAP, hxA, -⟩,
use A,
repeat { split } ; assumption,
end
-- 4ª demostración
example
(h : particion P)
: reflexive (relacion P) :=
begin
intro x,
rcases (h.1 x) with ⟨A, hAP, hxA, -⟩,
use [A, ⟨hAP, hxA, hxA⟩],
end
Se puede interactuar con la prueba anterior en esta sesión con Lean.
En los comentarios se pueden escribir otras soluciones, escribiendo el código entre una línea con <pre lang="lean"> y otra con </pre>
[/expand]
[expand title=»Soluciones con Isabelle/HOL»]
theory Las_particiones_definen_relaciones_reflexivas
imports Main
begin
definition relacion :: "('a set) set ⇒ 'a ⇒ 'a ⇒ bool" where
"relacion P x y ⟷ (∃A∈P. x ∈ A ∧ y ∈ A)"
definition particion :: "('a set) set ⇒ bool" where
"particion P ⟷ (∀x. (∃B∈P. x ∈ B ∧ (∀C∈P. x ∈ C ⟶ B = C))) ∧ {} ∉ P"
ru(* 1ª demostración *)
lemma
assumes "particion P"
shows "reflp (relacion P)"
proof (rule reflpI)
fix x
have "(∀x. (∃B∈P. x ∈ B ∧ (∀C∈P. x ∈ C ⟶ B = C))) ∧ {} ∉ P"
using assms by (unfold particion_def)
then have "∀x. (∃B∈P. x ∈ B ∧ (∀C∈P. x ∈ C ⟶ B = C))"
by (rule conjunct1)
then have "∃B∈P. x ∈ B ∧ (∀C∈P. x ∈ C ⟶ B = C)"
by (rule allE)
then obtain B where "B ∈ P ∧ (x ∈ B ∧ (∀C∈P. x ∈ C ⟶ B = C))"
by (rule someI2_bex)
then obtain B where "(B ∈ P ∧ x ∈ B) ∧ (∀C∈P. x ∈ C ⟶ B = C)"
by (simp only: conj_assoc)
then have "B ∈ P ∧ x ∈ B"
by (rule conjunct1)
then have "x ∈ B"
by (rule conjunct2)
then have "x ∈ B ∧ x ∈ B"
using ‹x ∈ B› by (rule conjI)
moreover have "B ∈ P"
using ‹B ∈ P ∧ x ∈ B› by (rule conjunct1)
ultimately have "∃B∈P. x ∈ B ∧ x ∈ B"
by (rule bexI)
then show "relacion P x x"
by (unfold relacion_def)
qed
(* 2ª demostración *)
lemma
assumes "particion P"
shows "reflp (relacion P)"
proof (rule reflpI)
fix x
obtain A where "A ∈ P ∧ x ∈ A"
using assms particion_def
by metis
then show "relacion P x x"
using relacion_def
by metis
qed
(* 3ª demostración *)
lemma
assumes "particion P"
shows "reflp (relacion P)"
using assms particion_def relacion_def
by (metis reflp_def)
end |
theory Las_particiones_definen_relaciones_reflexivas
imports Main
begin
definition relacion :: "('a set) set ⇒ 'a ⇒ 'a ⇒ bool" where
"relacion P x y ⟷ (∃A∈P. x ∈ A ∧ y ∈ A)"
definition particion :: "('a set) set ⇒ bool" where
"particion P ⟷ (∀x. (∃B∈P. x ∈ B ∧ (∀C∈P. x ∈ C ⟶ B = C))) ∧ {} ∉ P"
ru(* 1ª demostración *)
lemma
assumes "particion P"
shows "reflp (relacion P)"
proof (rule reflpI)
fix x
have "(∀x. (∃B∈P. x ∈ B ∧ (∀C∈P. x ∈ C ⟶ B = C))) ∧ {} ∉ P"
using assms by (unfold particion_def)
then have "∀x. (∃B∈P. x ∈ B ∧ (∀C∈P. x ∈ C ⟶ B = C))"
by (rule conjunct1)
then have "∃B∈P. x ∈ B ∧ (∀C∈P. x ∈ C ⟶ B = C)"
by (rule allE)
then obtain B where "B ∈ P ∧ (x ∈ B ∧ (∀C∈P. x ∈ C ⟶ B = C))"
by (rule someI2_bex)
then obtain B where "(B ∈ P ∧ x ∈ B) ∧ (∀C∈P. x ∈ C ⟶ B = C)"
by (simp only: conj_assoc)
then have "B ∈ P ∧ x ∈ B"
by (rule conjunct1)
then have "x ∈ B"
by (rule conjunct2)
then have "x ∈ B ∧ x ∈ B"
using ‹x ∈ B› by (rule conjI)
moreover have "B ∈ P"
using ‹B ∈ P ∧ x ∈ B› by (rule conjunct1)
ultimately have "∃B∈P. x ∈ B ∧ x ∈ B"
by (rule bexI)
then show "relacion P x x"
by (unfold relacion_def)
qed
(* 2ª demostración *)
lemma
assumes "particion P"
shows "reflp (relacion P)"
proof (rule reflpI)
fix x
obtain A where "A ∈ P ∧ x ∈ A"
using assms particion_def
by metis
then show "relacion P x x"
using relacion_def
by metis
qed
(* 3ª demostración *)
lemma
assumes "particion P"
shows "reflp (relacion P)"
using assms particion_def relacion_def
by (metis reflp_def)
end
En los comentarios se pueden escribir otras soluciones, escribiendo el código entre una línea con <pre lang="isar"> y otra con </pre>
[/expand]