]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/algebra/semigroups.ma
we removed some coercion detours and we added some coercions we really would like...
[helm.git] / helm / software / matita / library / algebra / semigroups.ma
index 7d68dbb360f9cdcef6dd4fdf3325cdd32fbc2b62..18d73b619402ecee81343ec6e9a9420181efa96f 100644 (file)
 
 include "higher_order_defs/functions.ma".
 
-(* Magmas *)
+(* Semigroups *)
 
-record Magma : Type≝
+record PreSemiGroup : Type≝
  { carrier:> Type;
    op: carrier → carrier → carrier
  }.
 
-interpretation "magma operation" 'middot a b = (op ? a b).
-
-(* Semigroups *)
+interpretation "Semigroup operation" 'middot a b = (op ? a b).
 
-record isSemiGroup (M:Magma) : Prop
- { op_associative: associative ? (op M) }.
+record IsSemiGroup (S:PreSemiGroup) : Prop 
+ { op_is_associative: associative ? (op S) }.
 
 record SemiGroup : Type≝
- { magma:> Magma;
-   semigroup_properties:> isSemiGroup magma
+ { pre_semi_group:> PreSemiGroup;
+   is_semi_group :> IsSemiGroup pre_semi_group
  }.
+
 definition is_left_unit ≝
- λS:SemiGroup. λe:S. ∀x:S. e·x = x.
+ λS:PreSemiGroup. λe:S. ∀x:S. e·x = x.
+
 definition is_right_unit ≝
- λS:SemiGroup. λe:S. ∀x:S. x·e = x.
+ λS:PreSemiGroup. λe:S. ∀x:S. x·e = x.
 
 theorem is_left_unit_to_is_right_unit_to_eq:
- ∀S:SemiGroup. ∀e,e':S.
+ ∀S:PreSemiGroup. ∀e,e':S.
   is_left_unit ? e → is_right_unit ? e' → e=e'.
  intros;
  rewrite < (H e');