X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Falgebra%2Fgroups.ma;h=bfb639af78deabee531767ff1d4a39b7245e763b;hb=e78cf74f8976cf0ca554f64baa9979d0423ee927;hp=ea345bc5469356224e3f102453264e0c3fa34542;hpb=e2a70d081915c180bca4ca9c9ea16971aa781db7;p=helm.git diff --git a/helm/software/matita/library/algebra/groups.ma b/helm/software/matita/library/algebra/groups.ma index ea345bc54..bfb639af7 100644 --- a/helm/software/matita/library/algebra/groups.ma +++ b/helm/software/matita/library/algebra/groups.ma @@ -12,8 +12,6 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/algebra/groups/". - include "algebra/monoids.ma". include "nat/le_arith.ma". include "datatypes/bool.ma". @@ -35,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. @@ -90,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); @@ -98,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); @@ -142,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˜ 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); @@ -164,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); @@ -189,24 +177,22 @@ notation "hvbox(x \sub H)" with precedence 79 for @{ 'subgroupimage $H $x }. interpretation "Subgroup image" 'subgroupimage H x = - (cic:/matita/algebra/groups/image.con _ _ - (cic:/matita/algebra/groups/morphism_of_subgroup.con _ H) x). + (image ?? (morphism_OF_subgroup ? H) x). definition member_of_subgroup ≝ λG.λH:subgroup G.λx:G.∃y.x=y \sub H. -notation "hvbox(x break ∈ H)" with precedence 79 +notation "hvbox(x break \in H)" with precedence 79 for @{ 'member_of $x $H }. -notation "hvbox(x break ∉ H)" with precedence 79 +notation "hvbox(x break \notin H)" with precedence 79 for @{ 'not_member_of $x $H }. interpretation "Member of subgroup" 'member_of x H = - (cic:/matita/algebra/groups/member_of_subgroup.con _ H x). + (member_of_subgroup ? H x). interpretation "Not member of subgroup" 'not_member_of x H = - (cic:/matita/logic/connectives/Not.con - (cic:/matita/algebra/groups/member_of_subgroup.con _ H x)). + (Not (member_of_subgroup ? H x)). (* Left cosets *) @@ -217,32 +203,31 @@ 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 = - (cic:/matita/algebra/groups/left_coset.ind#xpointer(1/1/1) _ 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 = - (cic:/matita/algebra/groups/member_of_left_coset.con _ C x). + (member_of_left_coset ? C x). definition left_coset_eq ≝ λG.λC,C':left_coset G. ∀x.((element ? C)·x \sub (subgrp ? C)) ∈ C'. -interpretation "Left cosets equality" 'eq C C' = - (cic:/matita/algebra/groups/left_coset_eq.con _ C C'). +interpretation "Left cosets equality" 'eq t C C' = (left_coset_eq t C C'). definition left_coset_disjoint ≝ λG.λC,C':left_coset G. ∀x.¬(((element ? C)·x \sub (subgrp ? C)) ∈ C'). -notation "hvbox(a break ∥ b)" +notation "hvbox(a break \par b)" non associative with precedence 45 for @{ 'disjoint $a $b }. interpretation "Left cosets disjoint" 'disjoint C C' = - (cic:/matita/algebra/groups/left_coset_disjoint.con _ C C'). + (left_coset_disjoint ? C C'). (* The following should be a one-shot alias! *) alias symbol "member_of" (instance 0) = "Member of subgroup". @@ -279,7 +264,7 @@ 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 < H3; @@ -296,10 +281,10 @@ qed. (*CSC: here the coercion Type_of_Group cannot be omitted. Why? *) theorem in_x_mk_left_coset_x_H: - ∀G.∀x:Type_of_Group G.∀H:subgroup G.x ∈ (x*H). + ∀G.∀x:Type_OF_Group G.∀H:subgroup G.x ∈ (x*H). intros; simplify; -apply (ex_intro ? ? 1); +apply (ex_intro ? ? ⅇ); rewrite > morphism_to_eq_f_1_1; rewrite > (e_is_right_unit ? G); reflexivity. @@ -316,10 +301,11 @@ record normal_subgroup (G:Group) : Type ≝ theorem foo: ∀G.∀H:normal_subgroup G.∀x.x*H=H*x. *) - +(* theorem member_of_left_coset_mk_left_coset_x_H_a_to_member_of_left_coset_mk_left_coset_y_H_b_to_member_of_left_coset_mk_left_coset_op_x_y_H_op_a_b: ∀G.∀H:normal_subgroup G.∀x,y,a,b. a ∈ (x*H) → b ∈ (y*H) → (a·b) ∈ ((x·y)*H). intros; simplify; -qed. \ No newline at end of file +qed. +*)