]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/algebra/groups.ma
we removed some coercion detours and we added some coercions we really would like...
[helm.git] / helm / software / matita / library / algebra / groups.ma
index bfb639af78deabee531767ff1d4a39b7245e763b..fd08a95dacb504e2cee3bfdf1a06948175c993f7 100644 (file)
@@ -18,22 +18,27 @@ include "datatypes/bool.ma".
 include "nat/compare.ma".
 
 record PreGroup : Type ≝
- { premonoid:> PreMonoid;
-   inv: premonoid -> premonoid
+ { pre_monoid:> PreMonoid;
+   inv: pre_monoid -> pre_monoid
  }.
 
+interpretation "Group inverse" 'invert x = (inv ? x).
+
 record isGroup (G:PreGroup) : Prop ≝
- { is_monoid:> isMonoid G;
-   inv_is_left_inverse: is_left_inverse (mk_Monoid ? is_monoid) (inv G);
-   inv_is_right_inverse: is_right_inverse (mk_Monoid ? is_monoid) (inv G)
+ { is_monoid           :> IsMonoid G;
+   inv_is_left_inverse :  is_left_inverse G (inv G);
+   inv_is_right_inverse:  is_right_inverse G (inv G)
  }.
+
 record Group : Type ≝
- { pregroup:> PreGroup;
-   group_properties:> isGroup pregroup
+ { pre_group:> PreGroup;
+   is_group:> isGroup pre_group
  }.
 
-interpretation "Group inverse" 'invert x = (inv ? x).
+definition Monoid_of_Group: Group → Monoid ≝
+ λG. mk_Monoid ? (is_group G).
+
+coercion Monoid_of_Group nocomposites.
 
 definition left_cancellable ≝
  λT:Type. λop: T -> T -> T.
@@ -52,25 +57,23 @@ intros (x y z);
 rewrite < (e_is_left_unit ? G);
 rewrite < (e_is_left_unit ? G z);
 rewrite < (inv_is_left_inverse ? G x);
-rewrite > (op_associative ? G);
-rewrite > (op_associative ? G);
+rewrite > (op_is_associative ? G);
+rewrite > (op_is_associative ? G);
 apply eq_f;
 assumption.
 qed.
 
-
 theorem eq_op_x_y_op_z_y_to_eq:
  ∀G:Group. right_cancellable G (op G).
 intros;
 unfold right_cancellable;
 unfold injective;
-simplify;fold simplify (op G); 
 intros (x y z);
 rewrite < (e_is_right_unit ? G);
 rewrite < (e_is_right_unit ? G z);
 rewrite < (inv_is_right_inverse ? G x);
-rewrite < (op_associative ? G);
-rewrite < (op_associative ? G);
+rewrite < (op_is_associative ? G);
+rewrite < (op_is_associative ? G);
 rewrite > H;
 reflexivity.
 qed.
@@ -104,7 +107,7 @@ theorem eq_opxy_z_to_eq_x_opzinvy:
  ∀G:Group. ∀x,y,z:G. x·y=z → x = z·y \sup -1.
 intros;
 apply (eq_op_x_y_op_z_y_to_eq ? y);
-rewrite > (op_associative ? G);
+rewrite > (op_is_associative ? G);
 rewrite > (inv_is_left_inverse ? G);
 rewrite > (e_is_right_unit ? G);
 assumption.
@@ -114,7 +117,7 @@ theorem eq_opxy_z_to_eq_y_opinvxz:
  ∀G:Group. ∀x,y,z:G. x·y=z → y = x \sup -1·z.
 intros;
 apply (eq_op_x_y_op_x_z_to_eq ? x);
-rewrite < (op_associative ? G);
+rewrite < (op_is_associative ? G);
 rewrite > (inv_is_right_inverse ? G);
 rewrite > (e_is_left_unit ? G);
 assumption.
@@ -125,8 +128,8 @@ theorem eq_inv_op_x_y_op_inv_y_inv_x:
 intros;
 apply (eq_op_x_y_op_z_y_to_eq ? (x·y));
 rewrite > (inv_is_left_inverse ? G);
-rewrite < (op_associative ? G);
-rewrite > (op_associative ? G (y \sup -1));
+rewrite < (op_is_associative ? G);
+rewrite > (op_is_associative ? G (y \sup -1));
 rewrite > (inv_is_left_inverse ? G);
 rewrite > (e_is_right_unit ? G);
 rewrite > (inv_is_left_inverse ? G);
@@ -246,8 +249,8 @@ exists;
   rewrite < H2;
   rewrite > eq_inv_op_x_y_op_inv_y_inv_x;
   rewrite > eq_inv_inv_x_x;
-  rewrite < (op_associative ? G);
-  rewrite < (op_associative ? G);
+  rewrite < (op_is_associative ? G);
+  rewrite < (op_is_associative ? G);
   rewrite > (inv_is_right_inverse ? G);
   rewrite > (e_is_left_unit ? G);
   reflexivity
@@ -266,12 +269,12 @@ elim H2;
 apply (ex_intro ? ? (x'·a \sup -1));
 rewrite > f_morph; 
 apply (eq_op_x_y_op_z_y_to_eq ? (a \sub H));
-rewrite > (op_associative ? G);
+rewrite > (op_is_associative ? G);
 rewrite < H3;
-rewrite > (op_associative ? G);
+rewrite > (op_is_associative ? G);
 rewrite < f_morph;
 rewrite > (inv_is_left_inverse ? H);
-rewrite < (op_associative ? G);
+rewrite < (op_is_associative ? G);
 rewrite > (inv_is_left_inverse ? G);
 rewrite > (e_is_left_unit ? G);
 rewrite < (f_morph ? ? H);