-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathLa_equipotencia_es_una_relacion_transitiva.thy
51 lines (42 loc) · 1.65 KB
/
La_equipotencia_es_una_relacion_transitiva.thy
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
(* La_equipotencia_es_una_relacion_transitiva.lean
-- La equipotencia es una relación transitiva.
-- José A. Alonso Jiménez <https://jaalonso.github.io>
-- Sevilla, 24-junio-2024
-- ------------------------------------------------------------------ *)
(* ---------------------------------------------------------------------
-- Dos conjuntos A y B son equipotentes (y se denota por A \<simeq> B) si
-- existe una aplicación biyectiva entre ellos. La equipotencia está
-- definida en Isabelle por
-- definition eqpoll :: "'a set \<Rightarrow> 'b set \<Rightarrow> bool" (infixl "\<approx>" 50)
-- where "eqpoll A B \<equiv> \<exists>f. bij_betw f A B"
--
-- Demostrar que la relación de equipotencia es transitiva.
-- ------------------------------------------------------------------ *)
theory La_equipotencia_es_una_relacion_transitiva
imports Main "HOL-Library.Equipollence"
begin
(* 1\<ordfeminine> demostración *)
lemma "transp (\<approx>)"
proof (rule transpI)
fix x y z :: "'a set"
assume "x \<approx> y" and "y \<approx> z"
show "x \<approx> z"
proof (unfold eqpoll_def)
obtain f where hf : "bij_betw f x y"
using \<open>x \<approx> y\<close> eqpoll_def by auto
obtain g where hg : "bij_betw g y z"
using \<open>y \<approx> z\<close> eqpoll_def by auto
have "bij_betw (g \<circ> f) x z"
using hf hg by (rule bij_betw_trans)
then show "\<exists>h. bij_betw h x z"
by auto
qed
qed
(* 2\<ordfeminine> demostración *)
lemma "transp (\<approx>)"
unfolding eqpoll_def transp_def
by (meson bij_betw_trans)
(* 3\<ordfeminine> demostración *)
lemma "transp (\<approx>)"
by (simp add: eqpoll_trans transpI)
end