-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathUnicidad_del_elemento_neutro_en_los_grupos.thy
50 lines (39 loc) · 1.24 KB
/
Unicidad_del_elemento_neutro_en_los_grupos.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
(* Unicidad_del_elemento_neutro_en_los_grupos.thy
-- Unicidad del elemento neutro en los grupos
-- José A. Alonso Jiménez <https://jaalonso.github.io>
-- Sevilla, 10-mayo-2024
-- ------------------------------------------------------------------ *)
(* ---------------------------------------------------------------------
-- Demostrar que un grupo sólo posee un elemento neutro.
-- ------------------------------------------------------------------ *)
theory Unicidad_del_elemento_neutro_en_los_grupos
imports Main
begin
context group
begin
(* 1\<ordfeminine> demostración *)
lemma
assumes "\<forall> x. x \<^bold>* e = x"
shows "e = \<^bold>1"
proof -
have "e = \<^bold>1 \<^bold>* e" by (simp only: left_neutral)
also have "\<dots> = \<^bold>1" using assms by (rule allE)
finally show "e = \<^bold>1" by this
qed
(* 2\<ordfeminine> demostración *)
lemma
assumes "\<forall> x. x \<^bold>* e = x"
shows "e = \<^bold>1"
proof -
have "e = \<^bold>1 \<^bold>* e" by simp
also have "\<dots> = \<^bold>1" using assms by simp
finally show "e = \<^bold>1" .
qed
(* 3\<ordfeminine> demostración *)
lemma
assumes "\<forall> x. x \<^bold>* e = x"
shows "e = \<^bold>1"
using assms
by (metis left_neutral)
end
end