]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/library/algebra/semigroups.ma
Preparing for 0.5.9 release.
[helm.git] / helm / software / matita / library / algebra / semigroups.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "higher_order_defs/functions.ma".
16
17 (* Semigroups *)
18
19 record PreSemiGroup : Type≝
20  { carrier:> Type;
21    op: carrier → carrier → carrier
22  }.
23
24 interpretation "Semigroup operation" 'middot a b = (op ? a b).
25
26 record IsSemiGroup (S:PreSemiGroup) : Prop ≝
27  { op_is_associative: associative ? (op S) }.
28
29 record SemiGroup : Type≝
30  { pre_semi_group:> PreSemiGroup;
31    is_semi_group :> IsSemiGroup pre_semi_group
32  }.
33
34 definition is_left_unit ≝
35  λS:PreSemiGroup. λe:S. ∀x:S. e·x = x.
36
37 definition is_right_unit ≝
38  λS:PreSemiGroup. λe:S. ∀x:S. x·e = x.
39
40 theorem is_left_unit_to_is_right_unit_to_eq:
41  ∀S:PreSemiGroup. ∀e,e':S.
42   is_left_unit ? e → is_right_unit ? e' → e=e'.
43  intros;
44  rewrite < (H e');
45  rewrite < (H1 e) in \vdash (? ? % ?).
46  reflexivity.
47 qed.