]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/library/algebra/semigroups.ma
- ok pr0 pr1 pr2
[helm.git] / matita / library / algebra / semigroups.ma
index 5b461d1a4f43edba6acc9e4dd448f0eceb157243..539228c8cdef6d600e5741397ce5d3a4020d4275 100644 (file)
@@ -18,15 +18,11 @@ include "higher_order_defs/functions.ma".
 
 (* Magmas *)
 
-record Magma : Type 
+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 }.
@@ -36,18 +32,14 @@ interpretation "magma operation" 'magma_op a b =
 
 (* Semigroups *)
 
-record isSemiGroup (M:Magma) : Prop 
- { associative: associative ? (op M) }.
+record isSemiGroup (M:Magma) : Prop≝
+ { op_associative: associative ? (op M) }.
 
-record SemiGroup : Type 
+record SemiGroup : Type≝
  { magma:> Magma;
    semigroup_properties:> isSemiGroup magma
  }.
  
-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. e·x = x.
  
@@ -59,6 +51,6 @@ theorem is_left_unit_to_is_right_unit_to_eq:
   is_left_unit ? e → is_right_unit ? e' → e=e'.
  intros;
  rewrite < (H e');
- rewrite < (H1 e) in \vdash (? ? % ?);
+ rewrite < (H1 e) in \vdash (? ? % ?).
  reflexivity.
 qed.