]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/algebra/groups.ma
more notation moved to core notation, unification of duplicated CProp connectives
[helm.git] / helm / software / matita / library / algebra / groups.ma
index aeb0d779ab98d46de7c0f24b1e6a37ca8b25eb44..29324c4aa748f9d4cf4f8b6d215c6d23f59bb2e0 100644 (file)
@@ -33,11 +33,7 @@ record Group : Type ≝
    group_properties:> isGroup pregroup
  }.
 
-notation "hvbox(x \sup (-1))" with precedence 89
-for @{ 'ginv $x }.
-
-interpretation "Group inverse" 'ginv x =
- (cic:/matita/algebra/groups/inv.con _ x).
+interpretation "Group inverse" 'invert x = (inv _ x).
 
 definition left_cancellable ≝
  λT:Type. λop: T -> T -> T.
@@ -88,7 +84,7 @@ reflexivity.
 qed.
 
 theorem eq_opxy_e_to_eq_x_invy:
- ∀G:Group. ∀x,y:G. x·y=1 → x=y \sup -1.
+ ∀G:Group. ∀x,y:G. x·y= → x=y \sup -1.
 intros;
 apply (eq_op_x_y_op_z_y_to_eq ? y);
 rewrite > (inv_is_left_inverse ? G);
@@ -96,7 +92,7 @@ assumption.
 qed.
 
 theorem eq_opxy_e_to_eq_invx_y:
- ∀G:Group. ∀x,y:G. x·y=1 → x \sup -1=y.
+ ∀G:Group. ∀x,y:G. x·y= → x \sup -1=y.
 intros;
 apply (eq_op_x_y_op_x_z_to_eq ? x);
 rewrite > (inv_is_right_inverse ? G);
@@ -140,20 +136,14 @@ qed.
 (* Morphisms *)
 
 record morphism (G,G':Group) : Type ≝
- { image: G → G';
+ { image:1> G → G';
    f_morph: ∀x,y:G.image(x·y) = image x · image y
  }.
  
-notation "hvbox(f\tilde x)" with precedence 79
-for @{ 'morimage $f $x }.
-
-interpretation "Morphism image" 'morimage f x =
- (cic:/matita/algebra/groups/image.con _ _ f x).
 theorem morphism_to_eq_f_1_1:
- ∀G,G'.∀f:morphism G G'.f˜1 = 1.
+ ∀G,G'.∀f:morphism G G'.f ⅇ  = ⅇ.
 intros;
-apply (eq_op_x_y_op_z_y_to_eq ? (f˜1));
+apply (eq_op_x_y_op_z_y_to_eq ? (f ⅇ));
 rewrite > (e_is_left_unit ? G');
 rewrite < f_morph;
 rewrite > (e_is_left_unit ? G);
@@ -162,9 +152,9 @@ qed.
  
 theorem eq_image_inv_inv_image:
  ∀G,G'.∀f:morphism G G'.
-  ∀x.f˜(x \sup -1) = (f˜x) \sup -1.
+  ∀x.f (x \sup -1) = (f x) \sup -1.
 intros;
-apply (eq_op_x_y_op_z_y_to_eq ? (f˜x));
+apply (eq_op_x_y_op_z_y_to_eq ? (f x));
 rewrite > (inv_is_left_inverse ? G');
 rewrite < f_morph;
 rewrite > (inv_is_left_inverse ? G);