]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/dama/divisible_group.ma
sandwitch theorem done
[helm.git] / helm / software / matita / dama / divisible_group.ma
index 5c7025a9c03d9b30a6efed52ebd1249cc3a4fdc2..7d2eeed454871b3b273036d137561e511b216544 100644 (file)
@@ -17,15 +17,15 @@ set "baseuri" "cic:/matita/divisible_group/".
 include "nat/orders.ma".
 include "group.ma".
 
-let rec pow (G : abelian_group) (x : G) (n : nat) on n : G ≝
-  match n with [ O ⇒ 0 | S m ⇒ x + pow ? x m].
+let rec gpow (G : abelian_group) (x : G) (n : nat) on n : G ≝
+  match n with [ O ⇒ 0 | S m ⇒ x + gpow ? x m].
   
-interpretation "abelian group pow" 'exp x n =
-  (cic:/matita/divisible_group/pow.con _ x n).
+interpretation "additive abelian group pow" 'times n x =
+  (cic:/matita/divisible_group/gpow.con _ x n).
   
 record dgroup : Type ≝ {
   dg_carr:> abelian_group;
-  dg_prop: ∀x:dg_carr.∀n:nat.∃y.pow ? y (S n) ≈ x
+  dg_prop: ∀x:dg_carr.∀n:nat.∃y.S n * y ≈ x
 }.
 
 lemma divide: ∀G:dgroup.G → nat → G.
@@ -36,15 +36,15 @@ interpretation "divisible group divide" 'divide x n =
   (cic:/matita/divisible_group/divide.con _ x n).
 
 lemma divide_divides: 
-  ∀G:dgroup.∀x:G.∀n. pow ? (x / n) (S n) ≈ x.
+  ∀G:dgroup.∀x:G.∀n. S n * (x / n) ≈ x.
 intro G; cases G; unfold divide; intros (x n); simplify;
 cases (f x n); simplify; exact H;
 qed.  
   
-lemma feq_pow: ∀G:dgroup.∀x,y:G.∀n.x≈y → pow ? x n ≈ pow ? y n.
+lemma feq_pow: ∀G:dgroup.∀x,y:G.∀n.x≈y → n * x ≈ n * y.
 intros (G x y n H); elim n; [apply eq_reflexive]
-simplify; apply (Eq≈ (x + (pow ? y n1)) H1);
-apply (Eq≈ (y+pow ? y n1) H (eq_reflexive ??));
+simplify; apply (Eq≈ (x + (n1 * y)) H1);
+apply (Eq≈ (y+n1*y) H (eq_reflexive ??));
 qed.
 
 lemma div1: ∀G:dgroup.∀x:G.x/O ≈ x.
@@ -53,3 +53,33 @@ cases (f x O); simplify; simplify in H; intro; apply H;
 apply (ap_rewl ???? (plus_comm ???));
 apply (ap_rewl ???w (zero_neutral ??)); assumption;
 qed.
+
+lemma appow_ap: ∀G:dgroup.∀x,y:G.∀n.S n * x # S n * y → x # y.
+intros 4 (G x y n); elim n; [2:
+  simplify in a;
+  cases (applus ????? a); [assumption]
+  apply f; assumption;]
+apply (plus_cancr_ap ??? 0); assumption;
+qed.
+
+lemma pluspow: ∀G:dgroup.∀x,y:G.∀n.n * (x+y) ≈ n * x + n * y.
+intros (G x y n); elim n; [
+  simplify; apply (Eq≈ 0 ? (zero_neutral ? 0)); apply eq_reflexive]
+simplify; apply eq_sym; apply (Eq≈ (x+y+(n1*x+n1*y))); [
+  apply (Eq≈ (x+(n1*x+(y+(n1*y))))); [
+    apply eq_sym; apply plus_assoc;]
+  apply (Eq≈ (x+((n1*x+y+(n1*y))))); [
+    apply feq_plusl; apply plus_assoc;]
+  apply (Eq≈ (x+(y+n1*x+n1*y))); [
+    apply feq_plusl; apply feq_plusr; apply plus_comm;] 
+  apply (Eq≈ (x+(y+(n1*x+n1*y)))); [
+    apply feq_plusl; apply eq_sym; apply plus_assoc;]
+  apply plus_assoc;]
+apply feq_plusl; apply eq_sym; assumption;
+qed. 
+
+lemma powzero: ∀G:dgroup.∀n.n*0 ≈ 0. [2: apply dg_carr; apply G]
+intros; elim n; [simplify; apply eq_reflexive]
+simplify; apply (Eq≈ ? (zero_neutral ??)); assumption; 
+qed.
+