X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Falgebra%2Fgroups.ma;h=6cb99481241e95a8c945902816d8832e5d8094c9;hb=ee3f8d6fa92b051394a2ff7c71c03ac33a05182b;hp=360e75c9f06d52c85e931e84f42eed4fa68bee76;hpb=5ce2b9b14c6b76021c780080e31b283a5e061d28;p=helm.git diff --git a/helm/software/matita/library/algebra/groups.ma b/helm/software/matita/library/algebra/groups.ma index 360e75c9f..6cb994812 100644 --- a/helm/software/matita/library/algebra/groups.ma +++ b/helm/software/matita/library/algebra/groups.ma @@ -294,6 +294,7 @@ rewrite > (e_is_right_unit ? H); reflexivity. 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). intros; @@ -302,4 +303,24 @@ apply (ex_intro ? ? 1); rewrite > morphism_to_eq_f_1_1; rewrite > (e_is_right_unit ? G); reflexivity. -qed. \ No newline at end of file +qed. + +(* Normal Subgroups *) + +record normal_subgroup (G:Group) : Type ≝ + { ns_subgroup:> subgroup G; + normal:> ∀x:G.∀y:ns_subgroup.(x·y \sub ns_subgroup·x \sup -1) ∈ ns_subgroup + }. + +(*CSC: I have not defined yet right cosets +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. +*)