]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/algebra/semigroups.ma
added links to svn tarballs
[helm.git] / helm / matita / library / algebra / semigroups.ma
index bea3ae584a3c456af650c82b443628f64a575a38..5b461d1a4f43edba6acc9e4dd448f0eceb157243 100644 (file)
@@ -16,28 +16,49 @@ set "baseuri" "cic:/matita/algebra/semigroups".
 
 include "higher_order_defs/functions.ma".
 
-definition isSemiGroup ≝
- λC:Type. λop: C → C → C.associative C op.
+(* Magmas *)
+
+record Magma : Type ≝
+ { carrier:> Type;
+   op: carrier → carrier → carrier
+ }.
+
+notation < "M" for @{ 'carrier $M }.
+interpretation "carrier coercion" 'carrier S =
+ (cic:/matita/algebra/semigroups/carrier.con S).
+
+notation "hvbox(a break \middot b)" 
+  left associative with precedence 55
+for @{ 'magma_op $a $b }.
+
+interpretation "magma operation" 'magma_op a b =
+ (cic:/matita/algebra/semigroups/op.con _ a b).
+
+(* Semigroups *)
+
+record isSemiGroup (M:Magma) : Prop ≝
+ { associative: associative ? (op M) }.
 
 record SemiGroup : Type ≝
- { carrier: Type;
-   op: carrier → carrier → carrier;
-   properties: isSemiGroup carrier op
+ { magma:> Magma;
+   semigroup_properties:> isSemiGroup magma
  }.
  
-coercion cic:/matita/algebra/semigroups/carrier.con.
-
+notation < "S" for @{ 'magma $S }.
+interpretation "magma coercion" 'magma S =
+ (cic:/matita/algebra/semigroups/magma.con S).
 definition is_left_unit ≝
- λS:SemiGroup. λe:S. ∀x:S. op S e x = x.
+ λS:SemiGroup. λe:S. ∀x:S. x = x.
  
 definition is_right_unit ≝
- λS:SemiGroup. λe:S. ∀x:S. op S x e = x.
+ λS:SemiGroup. λe:S. ∀x:S. e = x.
 
 theorem is_left_unit_to_is_right_unit_to_eq:
- ∀S:SemiGroup. ∀e1,e2:S.
-  is_left_unit ? e1 → is_right_unit ? e2 → e1=e2.
+ ∀S:SemiGroup. ∀e,e':S.
+  is_left_unit ? e → is_right_unit ? e' → e=e'.
  intros;
- rewrite < (H e2);
- rewrite < (H1 e1) in \vdash (? ? % ?);
+ rewrite < (H e');
+ rewrite < (H1 e) in \vdash (? ? % ?);
  reflexivity.
-qed.
\ No newline at end of file
+qed.