X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Falgebra%2Fgroups.ma;h=fd08a95dacb504e2cee3bfdf1a06948175c993f7;hb=647b419e96770d90a82d7a9e5e8843566a9f93ee;hp=f72711aaab9d33880f00caf8b312d6a780e79121;hpb=4dc47c9675ffd5fa50296ffaa9b5997501518c98;p=helm.git diff --git a/helm/software/matita/library/algebra/groups.ma b/helm/software/matita/library/algebra/groups.ma index f72711aaa..fd08a95da 100644 --- a/helm/software/matita/library/algebra/groups.ma +++ b/helm/software/matita/library/algebra/groups.ma @@ -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); @@ -177,7 +180,7 @@ notation "hvbox(x \sub H)" with precedence 79 for @{ 'subgroupimage $H $x }. interpretation "Subgroup image" 'subgroupimage H x = - (image _ _ (morphism_OF_subgroup _ H) x). + (image ?? (morphism_OF_subgroup ? H) x). definition member_of_subgroup ≝ λG.λH:subgroup G.λx:G.∃y.x=y \sub H. @@ -189,10 +192,10 @@ notation "hvbox(x break \notin H)" with precedence 79 for @{ 'not_member_of $x $H }. interpretation "Member of subgroup" 'member_of x H = - (member_of_subgroup _ H x). + (member_of_subgroup ? H x). interpretation "Not member of subgroup" 'not_member_of x H = - (Not (member_of_subgroup _ H x)). + (Not (member_of_subgroup ? H x)). (* Left cosets *) @@ -203,14 +206,14 @@ record left_coset (G:Group) : Type ≝ (* Here I would prefer 'magma_op, but this breaks something in the next definition *) interpretation "Left_coset" 'times x C = - (mk_left_coset _ x C). + (mk_left_coset ? x C). definition member_of_left_coset ≝ λG:Group.λC:left_coset G.λx:G. ∃y.x=(element ? C)·y \sub (subgrp ? C). interpretation "Member of left_coset" 'member_of x C = - (member_of_left_coset _ C x). + (member_of_left_coset ? C x). definition left_coset_eq ≝ λG.λC,C':left_coset G. @@ -227,7 +230,7 @@ notation "hvbox(a break \par b)" for @{ 'disjoint $a $b }. interpretation "Left cosets disjoint" 'disjoint C C' = - (left_coset_disjoint _ C C'). + (left_coset_disjoint ? C C'). (* The following should be a one-shot alias! *) alias symbol "member_of" (instance 0) = "Member of subgroup". @@ -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 @@ -264,14 +267,14 @@ apply H1; unfold member_of_subgroup; elim H2; apply (ex_intro ? ? (x'·a \sup -1)); -rewrite > f_morph; +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);