]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/algebra/groups.ma
more work on dama
[helm.git] / helm / software / matita / library / algebra / groups.ma
index 884e0c582e67d434c2f058ba83dd0a7fa7de54b0..aeb0d779ab98d46de7c0f24b1e6a37ca8b25eb44 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".
@@ -146,7 +144,7 @@ record morphism (G,G':Group) : Type ≝
    f_morph: ∀x,y:G.image(x·y) = image x · image y
  }.
  
-notation "hvbox(f˜ x)" with precedence 79
+notation "hvbox(f\tilde x)" with precedence 79
 for @{ 'morimage $f $x }.
 
 interpretation "Morphism image" 'morimage f x =
@@ -195,10 +193,10 @@ interpretation "Subgroup image" 'subgroupimage 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 =
@@ -237,7 +235,7 @@ 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 }.
 
@@ -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);