]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/library/algebra/groups.ma
added support for "polymorphic" coercions
[helm.git] / matita / library / algebra / groups.ma
index 5c5e1f00cf1d1f5823aa80269cdbda1298e1f6ab..7a675e3774fcec7e8638444d46db6ece963c542a 100644 (file)
@@ -25,7 +25,7 @@ record PreGroup : Type ≝
  }.
 
 record isGroup (G:PreGroup) : Prop ≝
- { is_monoid: isMonoid G;
+ { 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)
  }.
@@ -73,11 +73,11 @@ intros;
 unfold left_cancellable;
 unfold injective;
 intros (x y z);
-rewrite < (e_is_left_unit ? (is_monoid ? G));
-rewrite < (e_is_left_unit ? (is_monoid ? G) z);
+rewrite < (e_is_left_unit ? G);
+rewrite < (e_is_left_unit ? G z);
 rewrite < (inv_is_left_inverse ? G x);
-rewrite > (associative ? (is_semi_group ? (is_monoid ? G)));
-rewrite > (associative ? (is_semi_group ? (is_monoid ? G)));
+rewrite > (associative ? (is_semi_group ? ( G)));
+rewrite > (associative ? (is_semi_group ? ( G)));
 apply eq_f;
 assumption.
 qed.
@@ -90,11 +90,11 @@ unfold right_cancellable;
 unfold injective;
 simplify;fold simplify (op G); 
 intros (x y z);
-rewrite < (e_is_right_unit ? (is_monoid ? G));
-rewrite < (e_is_right_unit ? (is_monoid ? G) z);
+rewrite < (e_is_right_unit ? ( G));
+rewrite < (e_is_right_unit ? ( G) z);
 rewrite < (inv_is_right_inverse ? G x);
-rewrite < (associative ? (is_semi_group ? (is_monoid ? G)));
-rewrite < (associative ? (is_semi_group ? (is_monoid ? G)));
+rewrite < (associative ? (is_semi_group ? ( G)));
+rewrite < (associative ? (is_semi_group ? ( G)));
 rewrite > H;
 reflexivity.
 qed.
@@ -128,9 +128,9 @@ 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 > (associative ? (is_semi_group ? (is_monoid ? G)));
+rewrite > (associative ? G);
 rewrite > (inv_is_left_inverse ? G);
-rewrite > (e_is_right_unit ? (is_monoid ? G));
+rewrite > (e_is_right_unit ? G);
 assumption.
 qed.
 
@@ -138,7 +138,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 < (associative ? (is_semi_group ? (is_monoid ? G)));
+rewrite < (associative ? G);
 rewrite > (inv_is_right_inverse ? G);
 rewrite > (e_is_left_unit ? (is_monoid ? G));
 assumption.
@@ -161,9 +161,9 @@ theorem morphism_to_eq_f_1_1:
  ∀G,G'.∀f:morphism G G'.f˜1 = 1.
 intros;
 apply (eq_op_x_y_op_z_y_to_eq G' (f˜1));
-rewrite > (e_is_left_unit ? (is_monoid ? G') ?);
+rewrite > (e_is_left_unit ? G' ?);
 rewrite < (f_morph ? ? f);
-rewrite > (e_is_left_unit ? (is_monoid ? G));
+rewrite > (e_is_left_unit ? G);
 reflexivity.
 qed.
  
@@ -176,4 +176,134 @@ rewrite > (inv_is_left_inverse ? G');
 rewrite < (f_morph ? ? f);
 rewrite > (inv_is_left_inverse ? G);
 apply (morphism_to_eq_f_1_1 ? ? f).
-qed.
\ No newline at end of file
+qed.
+
+record monomorphism (G,G':Group) : Type ≝
+ { morphism: morphism G G';
+   injective: injective ? ? (image ? ? morphism)
+ }.
+
+(* Subgroups *)
+
+record subgroup (G:Group) : Type ≝
+ { group: Group;
+   embed: monomorphism group G
+ }.
+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.con _ _
+     (cic:/matita/algebra/groups/embed.con _ H))
+   x).
+
+definition belongs_to_subgroup ≝
+ λG.λH:subgroup G.λx:G.∃y.x=y \sub H.
+
+notation "hvbox(x ∈ H)" with precedence 79
+for @{ 'belongs_to $x $H }.
+
+interpretation "Belongs to subgroup" 'belongs_to x H =
+ (cic:/matita/algebra/groups/belongs_to_subgroup.con _ H x).
+
+(* Left cosets *)
+
+record left_coset (G:Group) : Type ≝
+ { element: G;
+   subgrp: subgroup G
+ }.
+
+(* 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).
+
+definition belongs_to_left_coset ≝
+ λG:Group.λC:left_coset G.λx:G.
+  ∃y.x=(element ? C)·y \sub (subgrp ? C).
+
+interpretation "Belongs to left_coset" 'belongs_to x C =
+ (cic:/matita/algebra/groups/belongs_to_left_coset.con _ 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').
+
+definition left_coset_disjoint ≝
+ λG.λC,C':left_coset G.
+  ∀x.¬(((element ? C)·x \sub (subgrp ? C)) ∈ C'). 
+
+notation "hvbox(a break ∥ 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').
+
+(*
+(* The following should be a one-shot alias! *)
+alias symbol "belongs_to" (instance 0) = "Belongs to subgroup".
+theorem foo:
+ ∀G.∀x,y:(Type_of_Group G).∀H:subgroup G.
+  (x \sup -1 ·y) ∈ H → (mk_left_coset ? x H) = (mk_left_coset ? y H).
+intros;
+unfold left_coset_eq;
+simplify in ⊢ (? → ? ? ? (? ? ? (? ? ? (? ? ? (? ? %)) ?)));
+simplify in ⊢ (? → ? ? ? (? ? % ?));
+simplify in ⊢ (? % → ?);
+intros;
+unfold belongs_to_left_coset;
+simplify in ⊢ (? ? (λy:?.? ? ? (? ? ? (? ? ? (? ? ? (? ? %)) ?))));
+simplify in ⊢ (? ? (λy:? %.?));
+simplify in ⊢ (? ? (λy:?.? ? ? (? ? % ?)));
+unfold belongs_to_subgroup in H1;
+elim H1;
+clear H1;
+exists;
+[apply ((a \sub H)\sup-1 · x1)
+| 
+].
+qed.
+*)
+
+(*theorem foo:
+ \forall G:Group. \forall x1,x2:G. \forall H:subgroup G.
+  x1*x2^-1 \nin H \to x1*H does_not_overlap x2*H
+
+theorem foo:
+ \forall x:G. \forall H:subgroup G. x \in x*H
+
+definition disjoinct
+ (T: Type) (n:nat) (S: \forall x:nat. x < n -> {S:Type * (S -> T)})
+:=
+ \forall i,j:nat. i < n \to j < n \to ...
+
+
+check
+ (λG.λH,H':left_coset G.λx:Type_of_Group (group ? (subgrp ? H)). (embed ? (subgrp ? H) x)).
+
+definition left_coset_eq ≝
+ λG.λH,H':left_coset G.
+  ∀x:group ? (subgrp ? H).
+   ex (group ? (subgroup ? H')) (λy.
+    (element ? H)·(embed ? (subgrp ? H) x) =
+    (element ? H')·(embed ? (subgrp ? H') y)).
+(*record left_coset (G:Group) : Type ≝
+ { subgroup: Group;
+   subgroup_is_subgroup: subgroup ≤ G;
+   element: G
+ }.
+
+definition left_coset_eq ≝
+ λG.λH,H':left_coset G.
+  ∀x:subgroup ? H.
+   ex (subgroup ? H') (λy.
+    (element ? H)·(embed ? ? (subgroup_is_subgroup ? H) ˜ x) =
+    (element ? H')·(embed ? ? (subgroup_is_subgroup ? H') ˜ y)).
+*)
+*)