]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/algebra/groups.ma
Most of the time, URIs can now be replaced with identifiers in "interpretation".
[helm.git] / helm / software / matita / library / algebra / groups.ma
index ea345bc5469356224e3f102453264e0c3fa34542..c83b3d2d0cc4359f982cf24cdcab9b869e72489f 100644 (file)
@@ -12,8 +12,6 @@
 (*                                                                        *)
 (**************************************************************************)
 
-set "baseuri" "cic:/matita/algebra/groups/".
-
 include "algebra/monoids.ma".
 include "nat/le_arith.ma".
 include "datatypes/bool.ma".
@@ -190,7 +188,7 @@ 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).
+   (cic:/matita/algebra/groups/morphism_OF_subgroup.con _ H) x).
 
 definition member_of_subgroup ≝
  λG.λH:subgroup G.λx:G.∃y.x=y \sub H.
@@ -296,7 +294,7 @@ 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);
@@ -316,10 +314,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.
+*)