aboutsummaryrefslogtreecommitdiff
path: root/coqprime/PrimalityTest
diff options
context:
space:
mode:
Diffstat (limited to 'coqprime/PrimalityTest')
-rw-r--r--coqprime/PrimalityTest/Cyclic.v244
-rw-r--r--coqprime/PrimalityTest/EGroup.v605
-rw-r--r--coqprime/PrimalityTest/Euler.v88
-rw-r--r--coqprime/PrimalityTest/FGroup.v123
-rw-r--r--coqprime/PrimalityTest/IGroup.v253
-rw-r--r--coqprime/PrimalityTest/Lagrange.v179
-rw-r--r--coqprime/PrimalityTest/LucasLehmer.v597
-rw-r--r--coqprime/PrimalityTest/Makefile.bak203
-rw-r--r--coqprime/PrimalityTest/Note.pdfbin0 -> 134038 bytes
-rw-r--r--coqprime/PrimalityTest/PGroup.v347
-rw-r--r--coqprime/PrimalityTest/Pepin.v123
-rw-r--r--coqprime/PrimalityTest/Pocklington.v261
-rw-r--r--coqprime/PrimalityTest/PocklingtonCertificat.v759
-rw-r--r--coqprime/PrimalityTest/Proth.v120
-rw-r--r--coqprime/PrimalityTest/Root.v239
-rw-r--r--coqprime/PrimalityTest/Zp.v411
16 files changed, 4552 insertions, 0 deletions
diff --git a/coqprime/PrimalityTest/Cyclic.v b/coqprime/PrimalityTest/Cyclic.v
new file mode 100644
index 000000000..c25f683ca
--- /dev/null
+++ b/coqprime/PrimalityTest/Cyclic.v
@@ -0,0 +1,244 @@
+
+(*************************************************************)
+(* This file is distributed under the terms of the *)
+(* GNU Lesser General Public License Version 2.1 *)
+(*************************************************************)
+(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
+(*************************************************************)
+
+(***********************************************************************
+ Cyclic.v
+
+ Proof that an abelien ring is cyclic
+ ************************************************************************)
+Require Import ZCAux.
+Require Import List.
+Require Import Root.
+Require Import UList.
+Require Import IGroup.
+Require Import EGroup.
+Require Import FGroup.
+
+Open Scope Z_scope.
+
+Section Cyclic.
+
+Variable A: Set.
+Variable plus mult: A -> A -> A.
+Variable op: A -> A.
+Variable zero one: A.
+Variable support: list A.
+Variable e: A.
+
+Hypothesis A_dec: forall a b: A, {a = b} + {a <> b}.
+Hypothesis e_not_zero: zero <> e.
+Hypothesis support_ulist: ulist support.
+Hypothesis e_in_support: In e support.
+Hypothesis zero_in_support: In zero support.
+Hypothesis mult_internal: forall a b, In a support -> In b support -> In (mult a b) support.
+Hypothesis mult_assoc: forall a b c, In a support -> In b support -> In c support -> mult a (mult b c) = mult (mult a b) c.
+Hypothesis e_is_zero_l: forall a, In a support -> mult e a = a.
+Hypothesis e_is_zero_r: forall a, In a support -> mult a e = a.
+Hypothesis plus_internal: forall a b, In a support -> In b support -> In (plus a b) support.
+Hypothesis plus_zero: forall a, In a support -> plus zero a = a.
+Hypothesis plus_comm: forall a b, In a support -> In b support -> plus a b = plus b a.
+Hypothesis plus_assoc: forall a b c, In a support -> In b support -> In c support -> plus a (plus b c) = plus (plus a b) c.
+Hypothesis mult_zero: forall a, In a support -> mult zero a = zero.
+Hypothesis mult_comm: forall a b, In a support -> In b support ->mult a b = mult b a.
+Hypothesis mult_plus_distr: forall a b c, In a support -> In b support -> In c support -> mult a (plus b c) = plus (mult a b) (mult a c).
+Hypothesis op_internal: forall a, In a support -> In (op a) support.
+Hypothesis plus_op_zero: forall a, In a support -> plus a (op a) = zero.
+Hypothesis mult_integral: forall a b, In a support -> In b support -> mult a b = zero -> a = zero \/ b = zero.
+
+Definition IA := (IGroup A mult support e A_dec support_ulist e_in_support mult_internal
+ mult_assoc
+ e_is_zero_l e_is_zero_r).
+
+Hint Resolve (fun x => isupport_incl _ mult support e A_dec x).
+
+Theorem gpow_evaln: forall n, 0 < n ->
+ exists p, (length p <= Zabs_nat n)%nat /\ (forall i, In i p -> In i support) /\
+ forall x, In x IA.(s) -> eval A plus mult zero (zero::p) x = gpow x IA n.
+intros n Hn; generalize Hn; pattern n; apply natlike_ind; auto with zarith.
+intros H1; contradict H1; auto with zarith.
+intros x Hx Rec _.
+case Zle_lt_or_eq with (1 := Hx); clear Hx; intros Hx; subst; simpl.
+case Rec; auto; simpl; intros p (Hp1, (Hp2, Hp3)); clear Rec.
+exists (zero::p); split; simpl.
+rewrite Zabs_nat_Zsucc; auto with arith zarith.
+split.
+intros i [Hi | Hi]; try rewrite <- Hi; auto.
+intros x1 Hx1; simpl.
+rewrite Hp3; repeat rewrite plus_zero; unfold Zsucc; try rewrite gpow_add; auto with zarith.
+rewrite gpow_1; try apply mult_comm; auto.
+apply (fun x => isupport_incl _ mult support e A_dec x); auto.
+change (In (gpow x1 IA x) IA.(s)).
+apply gpow_in; auto.
+apply mult_internal; auto.
+apply (fun x => isupport_incl _ mult support e A_dec x); auto.
+change (In (gpow x1 IA x) IA.(s)).
+apply gpow_in; auto.
+exists (e:: nil); split; simpl.
+compute; auto with arith.
+split.
+intros i [Hi | Hi]; try rewrite <- Hi; auto; case Hi.
+intros x Hx; simpl.
+rewrite plus_zero; rewrite (fun x => mult_comm x zero); try rewrite mult_zero; auto.
+rewrite plus_comm; try rewrite plus_zero; auto.
+Qed.
+
+Definition check_list_gpow: forall l n, (incl l IA.(s)) -> {forall a, In a l -> gpow a IA n = e} + {exists a, In a l /\ gpow a IA n <> e}.
+intros l n; elim l; simpl; auto.
+intros H; left; intros a H1; case H1.
+intros a l1 Rec H.
+case (A_dec (gpow a IA n) e); intros H2.
+case Rec; try intros H3.
+apply incl_tran with (2 := H); auto with datatypes.
+left; intros a1 H4; case H4; auto.
+intros H5; rewrite <- H5; auto.
+right; case H3; clear H3; intros a1 (H3, H4).
+exists a1; auto.
+right; exists a; auto.
+Defined.
+
+
+Theorem prime_power_div: forall p q i, prime p -> 0 <= q -> 0 <= i -> (q | p ^ i) -> exists j, 0 <= j <= i /\ q = p ^ j.
+intros p q i Hp Hq Hi H.
+assert (Hp1: 0 < p).
+apply Zlt_le_trans with 2; try apply prime_ge_2; auto with zarith.
+pattern q; apply prime_div_induction with (p ^ i); auto with zarith.
+exists 0; rewrite Zpower_0_r; auto with zarith.
+intros p1 i1 Hp2 Hi1 H1.
+case Zle_lt_or_eq with (1 := Hi1); clear Hi1; intros Hi1; subst.
+assert (Heq: p1 = p).
+apply prime_div_Zpower_prime with i; auto.
+apply Zdivide_trans with (2 := H1).
+apply Zpower_divide; auto with zarith.
+exists i1; split; auto; try split; auto with zarith.
+case (Zle_or_lt i1 i); auto; intros H2.
+absurd (p1 ^ i1 <= p ^ i).
+apply Zlt_not_le; rewrite Heq; apply Zpower_lt_monotone; auto with zarith.
+apply Zlt_le_trans with 2; try apply prime_ge_2; auto with zarith.
+apply Zdivide_le; auto with zarith.
+rewrite Heq; auto.
+exists 0; repeat rewrite Zpower_exp_0; auto with zarith.
+intros p1 q1 Hpq (j1,((Hj1, Hj2), Hj3)) (j2, ((Hj4, Hj5), Hj6)).
+case Zle_lt_or_eq with (1 := Hj1); clear Hj1; intros Hj1; subst.
+case Zle_lt_or_eq with (1 := Hj4); clear Hj4; intros Hj4; subst.
+inversion Hpq as [ H0 H1 H2].
+absurd (p | 1).
+intros H3; absurd (1 < p).
+apply Zle_not_lt; apply Zdivide_le; auto with zarith.
+apply Zlt_le_trans with 2; try apply prime_ge_2; auto with zarith.
+apply H2; apply Zpower_divide; auto with zarith.
+exists j1; rewrite Zpower_0_r; auto with zarith.
+exists j2; rewrite Zpower_0_r; auto with zarith.
+Qed.
+
+Theorem inj_lt_inv: forall n m : nat, Z_of_nat n < Z_of_nat m -> (n < m)%nat.
+intros n m H; case (le_or_lt m n); auto; intros H1; contradict H.
+apply Zle_not_lt; apply inj_le; auto.
+Qed.
+
+Theorem not_all_solutions: forall i, 0 < i < g_order IA -> exists a, In a IA.(s) /\ gpow a IA i <> e.
+intros i (Hi, Hi2).
+case (check_list_gpow IA.(s) i); try intros H; auto with datatypes.
+case (gpow_evaln i); auto; intros p (Hp1, (Hp2, Hp3)).
+absurd ((op e) = zero).
+intros H1; case e_not_zero.
+rewrite <- (plus_op_zero e); try rewrite H1; auto.
+rewrite plus_comm; auto.
+apply (root_max_is_zero _ (fun x => In x support) plus mult op zero) with (l := IA.(s)) (p := op e :: p); auto with datatypes.
+simpl; intros x [Hx | Hx]; try rewrite <- Hx; auto.
+intros x Hx.
+generalize (Hp3 _ Hx); simpl; rewrite plus_zero; auto.
+intros tmp; rewrite tmp; clear tmp.
+rewrite H; auto; rewrite plus_comm; auto with datatypes.
+apply mult_internal; auto.
+apply eval_P; auto.
+simpl; apply lt_le_S; apply le_lt_trans with (1 := Hp1).
+apply inj_lt_inv.
+rewrite inj_Zabs_nat; auto with zarith.
+rewrite Zabs_eq; auto with zarith.
+Qed.
+
+Theorem divide_g_order_e_order: forall n, 0 <= n -> (n | g_order IA) -> exists a, In a IA.(s) /\ e_order A_dec a IA = n.
+intros n Hn H.
+assert (Hg: 0 < g_order IA).
+apply g_order_pos.
+assert (He: forall a, 0 <= e_order A_dec a IA).
+intros a; apply Zlt_le_weak; apply e_order_pos.
+pattern n; apply prime_div_induction with (n := g_order IA); auto.
+exists e; split; auto.
+apply IA.(e_in_s).
+apply Zle_antisym.
+apply Zdivide_le; auto with zarith.
+apply e_order_divide_gpow; auto with zarith.
+apply IA.(e_in_s).
+rewrite gpow_1; auto.
+apply IA.(e_in_s).
+match goal with |- (_ <= ?X) => assert (0 < X) end; try apply e_order_pos; auto with zarith.
+intros p i Hp Hi K.
+assert (Hp1: 0 < p).
+apply Zlt_le_trans with 2; try apply prime_ge_2; auto with zarith.
+assert (Hi1: 0 < p ^ i).
+apply Zpower_gt_0; auto.
+case Zle_lt_or_eq with (1 := Hi); clear Hi; intros Hi; subst.
+case (not_all_solutions (g_order IA / p)).
+apply Zdivide_Zdiv_lt_pos; auto with zarith.
+apply Zlt_le_trans with 2; try apply prime_ge_2; auto with zarith.
+apply Zdivide_trans with (2 := K).
+apply Zpower_divide; auto.
+intros a (Ha1, Ha2).
+exists (gpow a IA (g_order IA / p ^ i)); split.
+apply gpow_in; auto.
+match goal with |- ?X = ?Y => assert (H1: (X | Y) ) end; auto.
+apply e_order_divide_gpow; auto with zarith.
+apply gpow_in; auto.
+rewrite <- gpow_gpow; auto with zarith.
+rewrite Zmult_comm; rewrite <- Zdivide_Zdiv_eq; auto with zarith.
+apply fermat_gen; auto.
+apply Z_div_pos; auto with zarith.
+case prime_power_div with (4 := H1); auto with zarith.
+intros j ((Hj1, Hj2), Hj3).
+case Zle_lt_or_eq with (1 := Hj2); intros Hj4; subst; auto.
+case Ha2.
+replace (g_order IA) with (((g_order IA / p ^i) * p ^ j) * p ^ (i - j - 1) * p).
+rewrite Z_div_mult; auto with zarith.
+repeat rewrite gpow_gpow; auto with zarith.
+rewrite <- Hj3.
+rewrite gpow_e_order_is_e; auto with zarith.
+rewrite gpow_e; auto.
+apply Zlt_le_weak; apply Zpower_gt_0; auto with zarith.
+apply gpow_in; auto.
+apply Z_div_pos; auto with zarith.
+apply Zmult_le_0_compat; try apply Z_div_pos; auto with zarith.
+pattern p at 4; rewrite <- Zpower_1_r.
+repeat rewrite <- Zmult_assoc; repeat rewrite <- Zpower_exp; auto with zarith.
+replace (j + (i - j - 1 + 1)) with i; auto with zarith.
+apply sym_equal; rewrite Zmult_comm; apply Zdivide_Zdiv_eq; auto with zarith.
+rewrite Zpower_0_r; exists e; split.
+apply IA.(e_in_s).
+match goal with |- ?X = 1 => assert (tmp: 0 < X); try apply e_order_pos;
+case Zle_lt_or_eq with 1 X; auto with zarith; clear tmp; intros H1 end.
+absurd (gpow IA.(FGroup.e) IA 1 = IA.(FGroup.e)).
+apply gpow_e_order_lt_is_not_e with A_dec; auto with zarith.
+apply gpow_e; auto with zarith.
+intros p q H1 (a, (Ha1, Ha2)) (b, (Hb1, Hb2)).
+exists (mult a b); split.
+apply IA.(internal); auto.
+rewrite <- Ha2; rewrite <- Hb2; apply order_mult; auto.
+rewrite Ha2; rewrite Hb2; auto.
+Qed.
+
+Set Implicit Arguments.
+Definition cyclic (A: Set) A_dec (op: A -> A -> A) (G: FGroup op):= exists a, In a G.(s) /\ e_order A_dec a G = g_order G.
+Unset Implicit Arguments.
+
+Theorem cyclic_field: cyclic A_dec IA.
+red; apply divide_g_order_e_order; auto.
+apply Zlt_le_weak; apply g_order_pos.
+exists 1; ring.
+Qed.
+
+End Cyclic.
diff --git a/coqprime/PrimalityTest/EGroup.v b/coqprime/PrimalityTest/EGroup.v
new file mode 100644
index 000000000..fd543fe04
--- /dev/null
+++ b/coqprime/PrimalityTest/EGroup.v
@@ -0,0 +1,605 @@
+
+(*************************************************************)
+(* This file is distributed under the terms of the *)
+(* GNU Lesser General Public License Version 2.1 *)
+(*************************************************************)
+(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
+(*************************************************************)
+
+(**********************************************************************
+ EGroup.v
+
+ Given an element a, create the group {e, a, a^2, ..., a^n}
+ **********************************************************************)
+Require Import ZArith.
+Require Import Tactic.
+Require Import List.
+Require Import ZCAux.
+Require Import ZArith Znumtheory.
+Require Import Wf_nat.
+Require Import UList.
+Require Import FGroup.
+Require Import Lagrange.
+
+Open Scope Z_scope.
+
+Section EGroup.
+
+Variable A: Set.
+
+Variable A_dec: forall a b: A, {a = b} + {~ a = b}.
+
+Variable op: A -> A -> A.
+
+Variable a: A.
+
+Variable G: FGroup op.
+
+Hypothesis a_in_G: In a G.(s).
+
+
+(**************************************
+ The power function for the group
+ **************************************)
+
+Set Implicit Arguments.
+Definition gpow n := match n with Zpos p => iter_pos p _ (op a) G.(e) | _ => G.(e) end.
+Unset Implicit Arguments.
+
+Theorem gpow_0: gpow 0 = G.(e).
+simpl; sauto.
+Qed.
+
+Theorem gpow_1 : gpow 1 = a.
+simpl; sauto.
+Qed.
+
+(**************************************
+ Some properties of the power function
+ **************************************)
+
+Theorem gpow_in: forall n, In (gpow n) G.(s).
+intros n; case n; simpl; auto.
+intros p; apply iter_pos_invariant with (Inv := fun x => In x G.(s)); auto.
+Qed.
+
+Theorem gpow_op: forall b p, In b G.(s) -> iter_pos p _ (op a) b = op (iter_pos p _ (op a) G.(e)) b.
+intros b p; generalize b; elim p; simpl; auto; clear b p.
+intros p Rec b Hb.
+assert (H: In (gpow (Zpos p)) G.(s)).
+apply gpow_in.
+rewrite (Rec b); try rewrite (fun x y => Rec (op x y)); try rewrite (fun x y => Rec (iter_pos p A x y)); auto.
+repeat rewrite G.(assoc); auto.
+intros p Rec b Hb.
+assert (H: In (gpow (Zpos p)) G.(s)).
+apply gpow_in.
+rewrite (Rec b); try rewrite (fun x y => Rec (op x y)); try rewrite (fun x y => Rec (iter_pos p A x y)); auto.
+repeat rewrite G.(assoc); auto.
+intros b H; rewrite e_is_zero_r; auto.
+Qed.
+
+Theorem gpow_add: forall n m, 0 <= n -> 0 <= m -> gpow (n + m) = op (gpow n) (gpow m).
+intros n; case n.
+intros m _ _; simpl; apply sym_equal; apply e_is_zero_l; apply gpow_in.
+2: intros p m H; contradict H; auto with zarith.
+intros p1 m; case m.
+intros _ _; simpl; apply sym_equal; apply e_is_zero_r.
+exact (gpow_in (Zpos p1)).
+2: intros p2 _ H; contradict H; auto with zarith.
+intros p2 _ _; simpl.
+rewrite iter_pos_plus; rewrite (fun x y => gpow_op (iter_pos p2 A x y)); auto.
+exact (gpow_in (Zpos p2)).
+Qed.
+
+Theorem gpow_1_more:
+ forall n, 0 < n -> gpow n = G.(e) -> forall m, 0 <= m -> exists p, 0 <= p < n /\ gpow m = gpow p.
+intros n H1 H2 m Hm; generalize Hm; pattern m; apply Z_lt_induction; auto with zarith; clear m Hm.
+intros m Rec Hm.
+case (Zle_or_lt n m); intros H3.
+case (Rec (m - n)); auto with zarith.
+intros p (H4,H5); exists p; split; auto.
+replace m with (n + (m - n)); auto with zarith.
+rewrite gpow_add; try rewrite H2; try rewrite H5; sauto; auto with zarith.
+generalize gpow_in; sauto.
+exists m; auto.
+Qed.
+
+Theorem gpow_i: forall n m, 0 <= n -> 0 <= m -> gpow n = gpow (n + m) -> gpow m = G.(e).
+intros n m H1 H2 H3; generalize gpow_in; intro PI.
+apply g_cancel_l with (g:= G) (a := gpow n); sauto.
+rewrite <- gpow_add; try rewrite <- H3; sauto.
+Qed.
+
+(**************************************
+ We build the support by iterating the power function
+ **************************************)
+
+Set Implicit Arguments.
+
+Fixpoint support_aux (b: A) (n: nat) {struct n}: list A :=
+b::let c := op a b in
+ match n with
+ O => nil |
+ (S n1) =>if A_dec c G.(e) then nil else support_aux c n1
+ end.
+
+Definition support := support_aux G.(e) (Zabs_nat (g_order G)).
+
+Unset Implicit Arguments.
+
+(**************************************
+ Some properties of the support that helps to prove that we have a group
+ **************************************)
+
+Theorem support_aux_gpow:
+ forall n m b, 0 <= m -> In b (support_aux (gpow m) n) ->
+ exists p, (0 <= p < length (support_aux (gpow m) n))%nat /\ b = gpow (m + Z_of_nat p).
+intros n; elim n; simpl.
+intros n1 b Hm [H1 | H1]; exists 0%nat; simpl; rewrite Zplus_0_r; auto; case H1.
+intros n1 Rec m b Hm [H1 | H1].
+exists 0%nat; simpl; rewrite Zplus_0_r; auto; auto with arith.
+generalize H1; case (A_dec (op a (gpow m)) G.(e)); clear H1; simpl; intros H1 H2.
+case H2.
+case (Rec (1 + m) b); auto with zarith.
+rewrite gpow_add; auto with zarith.
+rewrite gpow_1; auto.
+intros p (Hp1, Hp2); exists (S p); split; auto with zarith.
+rewrite <- gpow_1.
+rewrite <- gpow_add; auto with zarith.
+rewrite inj_S; rewrite Hp2; eq_tac; auto with zarith.
+Qed.
+
+Theorem gpow_support_aux_not_e:
+ forall n m p, 0 <= m -> m < p < m + Z_of_nat (length (support_aux (gpow m) n)) -> gpow p <> G.(e).
+intros n; elim n; simpl.
+intros m p Hm (H1, H2); contradict H2; auto with zarith.
+intros n1 Rec m p Hm; case (A_dec (op a (gpow m)) G.(e)); simpl.
+intros _ (H1, H2); contradict H2; auto with zarith.
+assert (tmp: forall p, Zpos (P_of_succ_nat p) = 1 + Z_of_nat p).
+intros p1; apply trans_equal with (Z_of_nat (S p1)); auto; rewrite inj_S; auto with zarith.
+rewrite tmp.
+intros H1 (H2, H3); case (Zle_lt_or_eq (1 + m) p); auto with zarith; intros H4; subst.
+apply (Rec (1 + m)); try split; auto with zarith.
+rewrite gpow_add; auto with zarith.
+rewrite gpow_1; auto with zarith.
+rewrite gpow_add; try rewrite gpow_1; auto with zarith.
+Qed.
+
+Theorem support_aux_not_e: forall n m b, 0 <= m -> In b (tail (support_aux (gpow m) n)) -> ~ b = G.(e).
+intros n; elim n; simpl.
+intros m b Hm H; case H.
+intros n1 Rec m b Hm; case (A_dec (op a (gpow m)) G.(e)); intros H1 H2; simpl; auto.
+assert (Hm1: 0 <= 1 + m); auto with zarith.
+generalize( Rec (1 + m) b Hm1) H2; case n1; auto; clear Hm1.
+intros _ [H3 | H3]; auto.
+contradict H1; subst; auto.
+rewrite gpow_add; simpl; try rewrite e_is_zero_r; auto with zarith.
+intros n2; case (A_dec (op a (op a (gpow m))) G.(e)); intros H3.
+intros _ [H4 | H4].
+contradict H1; subst; auto.
+case H4.
+intros H4 [H5 | H5]; subst; auto.
+Qed.
+
+Theorem support_aux_length_le: forall n a, (length (support_aux a n) <= n + 1)%nat.
+intros n; elim n; simpl; auto.
+intros n1 Rec a1; case (A_dec (op a a1) G.(e)); simpl; auto with arith.
+Qed.
+
+Theorem support_aux_length_le_is_e:
+ forall n m, 0 <= m -> (length (support_aux (gpow m) n) <= n)%nat ->
+ gpow (m + Z_of_nat (length (support_aux (gpow m) n))) = G.(e) .
+intros n; elim n; simpl; auto.
+intros m _ H1; contradict H1; auto with arith.
+intros n1 Rec m Hm; case (A_dec (op a (gpow m)) G.(e)); simpl; intros H1.
+intros H2; rewrite Zplus_comm; rewrite gpow_add; simpl; try rewrite e_is_zero_r; auto with zarith.
+assert (tmp: forall p, Zpos (P_of_succ_nat p) = 1 + Z_of_nat p).
+intros p1; apply trans_equal with (Z_of_nat (S p1)); auto; rewrite inj_S; auto with zarith.
+rewrite tmp; clear tmp.
+rewrite <- gpow_1.
+rewrite <- gpow_add; auto with zarith.
+rewrite Zplus_assoc; rewrite (Zplus_comm 1); intros H2; apply Rec; auto with zarith.
+Qed.
+
+Theorem support_aux_in:
+ forall n m p, 0 <= m -> (p < length (support_aux (gpow m) n))% nat ->
+ (In (gpow (m + Z_of_nat p)) (support_aux (gpow m) n)).
+intros n; elim n; simpl; auto; clear n.
+intros m p Hm H1; replace p with 0%nat.
+left; eq_tac; auto with zarith.
+generalize H1; case p; simpl; auto with arith.
+intros n H2; contradict H2; apply le_not_lt; auto with arith.
+intros n1 Rec m p Hm; case (A_dec (op a (gpow m)) G.(e)); simpl; intros H1 H2; auto.
+replace p with 0%nat.
+left; eq_tac; auto with zarith.
+generalize H2; case p; simpl; auto with arith.
+intros n H3; contradict H3; apply le_not_lt; auto with arith.
+generalize H2; case p; simpl; clear H2.
+rewrite Zplus_0_r; auto.
+intros n.
+assert (tmp: forall p, Zpos (P_of_succ_nat p) = 1 + Z_of_nat p).
+intros p1; apply trans_equal with (Z_of_nat (S p1)); auto; rewrite inj_S; auto with zarith.
+rewrite tmp; clear tmp.
+rewrite <- gpow_1; rewrite <- gpow_add; auto with zarith.
+rewrite Zplus_assoc; rewrite (Zplus_comm 1); intros H2; right; apply Rec; auto with zarith.
+Qed.
+
+Theorem support_aux_ulist:
+ forall n m, 0 <= m -> (forall p, 0 <= p < m -> gpow (1 + p) <> G.(e)) -> ulist (support_aux (gpow m) n).
+intros n; elim n; auto; clear n.
+intros m _ _; auto.
+simpl; apply ulist_cons; auto.
+intros n1 Rec m Hm H.
+simpl; case (A_dec (op a (gpow m)) G.(e)); auto.
+intros He; apply ulist_cons; auto.
+intros H1; case (support_aux_gpow n1 (1 + m) (gpow m)); auto with zarith.
+rewrite gpow_add; try rewrite gpow_1; auto with zarith.
+intros p (Hp1, Hp2).
+assert (H2: gpow (1 + Z_of_nat p) = G.(e)).
+apply gpow_i with m; auto with zarith.
+rewrite Hp2; eq_tac; auto with zarith.
+case (Zle_or_lt m (Z_of_nat p)); intros H3; auto.
+2: case (H (Z_of_nat p)); auto with zarith.
+case (support_aux_not_e (S n1) m (gpow (1 + Z_of_nat p))); auto.
+rewrite gpow_add; auto with zarith; simpl; rewrite e_is_zero_r; auto.
+case (A_dec (op a (gpow m)) G.(e)); auto.
+intros _; rewrite <- gpow_1; repeat rewrite <- gpow_add; auto with zarith.
+replace (1 + Z_of_nat p) with ((1 + m) + (Z_of_nat (p - Zabs_nat m))); auto with zarith.
+apply support_aux_in; auto with zarith.
+rewrite inj_minus1; auto with zarith.
+rewrite inj_Zabs_nat; auto with zarith.
+rewrite Zabs_eq; auto with zarith.
+apply inj_le_rev.
+rewrite inj_Zabs_nat; auto with zarith.
+rewrite Zabs_eq; auto with zarith.
+rewrite <- gpow_1; repeat rewrite <- gpow_add; auto with zarith.
+apply (Rec (1 + m)); auto with zarith.
+intros p H1; case (Zle_lt_or_eq p m); intros; subst; auto with zarith.
+rewrite gpow_add; auto with zarith.
+rewrite gpow_1; auto.
+Qed.
+
+Theorem support_gpow: forall b, (In b support) -> exists p, 0 <= p < Z_of_nat (length support) /\ b = gpow p.
+intros b H; case (support_aux_gpow (Zabs_nat (g_order G)) 0 b); auto with zarith.
+intros p ((H1, H2), H3); exists (Z_of_nat p); repeat split; auto with zarith.
+apply inj_lt; auto.
+Qed.
+
+Theorem support_incl_G: incl support G.(s).
+intros a1 H; case (support_gpow a1); auto; intros p (H1, H2); subst; apply gpow_in.
+Qed.
+
+Theorem gpow_support_not_e: forall p, 0 < p < Z_of_nat (length support) -> gpow p <> G.(e).
+intros p (H1, H2); apply gpow_support_aux_not_e with (m := 0) (n := length G.(s)); simpl;
+ try split; auto with zarith.
+rewrite <- (Zabs_nat_Z_of_nat (length G.(s))); auto.
+Qed.
+
+Theorem support_not_e: forall b, In b (tail support) -> ~ b = G.(e).
+intros b H; apply (support_aux_not_e (Zabs_nat (g_order G)) 0); auto with zarith.
+Qed.
+
+Theorem support_ulist: ulist support.
+apply (support_aux_ulist (Zabs_nat (g_order G)) 0); auto with zarith.
+Qed.
+
+Theorem support_in_e: In G.(e) support.
+unfold support; case (Zabs_nat (g_order G)); simpl; auto with zarith.
+Qed.
+
+Theorem gpow_length_support_is_e: gpow (Z_of_nat (length support)) = G.(e).
+apply (support_aux_length_le_is_e (Zabs_nat (g_order G)) 0); simpl; auto with zarith.
+unfold g_order; rewrite Zabs_nat_Z_of_nat; apply ulist_incl_length.
+rewrite <- (Zabs_nat_Z_of_nat (length G.(s))); auto.
+exact support_ulist.
+rewrite <- (Zabs_nat_Z_of_nat (length G.(s))); auto.
+exact support_incl_G.
+Qed.
+
+Theorem support_in: forall p, 0 <= p < Z_of_nat (length support) -> In (gpow p) support.
+intros p (H, H1); unfold support.
+rewrite <- (Zabs_eq p); auto with zarith.
+rewrite <- (inj_Zabs_nat p); auto.
+generalize (support_aux_in (Zabs_nat (g_order G)) 0); simpl; intros H2; apply H2; auto with zarith.
+rewrite <- (fun x => Zabs_nat_Z_of_nat (@length A x)); auto.
+apply Zabs_nat_lt; split; auto.
+Qed.
+
+Theorem support_internal: forall a b, In a support -> In b support -> In (op a b) support.
+intros a1 b1 H1 H2.
+case support_gpow with (1 := H1); auto; intros p1 ((H3, H4), H5); subst.
+case support_gpow with (1 := H2); auto; intros p2 ((H5, H6), H7); subst.
+rewrite <- gpow_add; auto with zarith.
+case gpow_1_more with (m:= p1 + p2) (2 := gpow_length_support_is_e); auto with zarith.
+intros p3 ((H8, H9), H10); rewrite H10; apply support_in; auto with zarith.
+Qed.
+
+Theorem support_i_internal: forall a, In a support -> In (G.(i) a) support.
+generalize gpow_in; intros Hp.
+intros a1 H1.
+case support_gpow with (1 := H1); auto.
+intros p1 ((H2, H3), H4); case Zle_lt_or_eq with (1 := H2); clear H2; intros H2; subst.
+2: rewrite gpow_0; rewrite i_e; apply support_in_e.
+replace (G.(i) (gpow p1)) with (gpow (Z_of_nat (length support - Zabs_nat p1))).
+apply support_in; auto with zarith.
+rewrite inj_minus1.
+rewrite inj_Zabs_nat; auto with zarith.
+rewrite Zabs_eq; auto with zarith.
+apply inj_le_rev; rewrite inj_Zabs_nat; auto with zarith.
+rewrite Zabs_eq; auto with zarith.
+apply g_cancel_l with (g:= G) (a := gpow p1); sauto.
+rewrite <- gpow_add; auto with zarith.
+replace (p1 + Z_of_nat (length support - Zabs_nat p1)) with (Z_of_nat (length support)).
+rewrite gpow_length_support_is_e; sauto.
+rewrite inj_minus1; auto with zarith.
+rewrite inj_Zabs_nat; auto with zarith.
+rewrite Zabs_eq; auto with zarith.
+apply inj_le_rev; rewrite inj_Zabs_nat; auto with zarith.
+rewrite Zabs_eq; auto with zarith.
+Qed.
+
+(**************************************
+ We are now ready to build the group
+ **************************************)
+
+Definition Gsupport: (FGroup op).
+generalize support_incl_G; unfold incl; intros Ho.
+apply mkGroup with support G.(e) G.(i); sauto.
+apply support_ulist.
+apply support_internal.
+intros a1 b1 c1 H1 H2 H3; apply G.(assoc); sauto.
+apply support_in_e.
+apply support_i_internal.
+Defined.
+
+(**************************************
+ Definition of the order of an element
+ **************************************)
+Set Implicit Arguments.
+
+Definition e_order := Z_of_nat (length support).
+
+Unset Implicit Arguments.
+
+(**************************************
+ Some properties of the order of an element
+ **************************************)
+
+Theorem gpow_e_order_is_e: gpow e_order = G.(e).
+apply (support_aux_length_le_is_e (Zabs_nat (g_order G)) 0); simpl; auto with zarith.
+unfold g_order; rewrite Zabs_nat_Z_of_nat; apply ulist_incl_length.
+rewrite <- (Zabs_nat_Z_of_nat (length G.(s))); auto.
+exact support_ulist.
+rewrite <- (Zabs_nat_Z_of_nat (length G.(s))); auto.
+exact support_incl_G.
+Qed.
+
+Theorem gpow_e_order_lt_is_not_e: forall n, 1 <= n < e_order -> gpow n <> G.(e).
+intros n (H1, H2); apply gpow_support_not_e; auto with zarith.
+Qed.
+
+Theorem e_order_divide_g_order: (e_order | g_order G).
+change ((g_order Gsupport) | g_order G).
+apply lagrange; auto.
+exact support_incl_G.
+Qed.
+
+Theorem e_order_pos: 0 < e_order.
+unfold e_order, support; case (Zabs_nat (g_order G)); simpl; auto with zarith.
+Qed.
+
+Theorem e_order_divide_gpow: forall n, 0 <= n -> gpow n = G.(e) -> (e_order | n).
+generalize gpow_in; intros Hp.
+generalize e_order_pos; intros Hp1.
+intros n Hn; generalize Hn; pattern n; apply Z_lt_induction; auto; clear n Hn.
+intros n Rec Hn H.
+case (Zle_or_lt e_order n); intros H1.
+case (Rec (n - e_order)); auto with zarith.
+apply g_cancel_l with (g:= G) (a := gpow e_order); sauto.
+rewrite G.(e_is_zero_r); auto with zarith.
+rewrite <- gpow_add; try (rewrite gpow_e_order_is_e; rewrite <- H; eq_tac); auto with zarith.
+intros k Hk; exists (1 + k).
+rewrite Zmult_plus_distr_l; rewrite <- Hk; auto with zarith.
+case (Zle_lt_or_eq 0 n); auto with arith; intros H2; subst.
+contradict H; apply support_not_e.
+generalize H1; unfold e_order, support.
+case (Zabs_nat (g_order G)); simpl; auto.
+intros H3; contradict H3; auto with zarith.
+intros n1; case (A_dec (op a G.(e)) G.(e)); simpl; intros _ H3.
+contradict H3; auto with zarith.
+generalize H3; clear H3.
+assert (tmp: forall p, Zpos (P_of_succ_nat p) = 1 + Z_of_nat p).
+intros p1; apply trans_equal with (Z_of_nat (S p1)); auto; rewrite inj_S; auto with zarith.
+rewrite tmp; clear tmp; intros H3.
+change (In (gpow n) (support_aux (gpow 1) n1)).
+replace n with (1 + Z_of_nat (Zabs_nat n - 1)).
+apply support_aux_in; auto with zarith.
+rewrite <- (fun x => Zabs_nat_Z_of_nat (@length A x)).
+replace (Zabs_nat n - 1)%nat with (Zabs_nat (n - 1)).
+apply Zabs_nat_lt; split; auto with zarith.
+rewrite G.(e_is_zero_r) in H3; try rewrite gpow_1; auto with zarith.
+apply inj_eq_rev; rewrite inj_Zabs_nat; auto with zarith.
+rewrite Zabs_eq; auto with zarith.
+rewrite inj_minus1; auto with zarith.
+rewrite inj_Zabs_nat; auto with zarith.
+rewrite Zabs_eq; auto with zarith.
+apply inj_le_rev; rewrite inj_Zabs_nat; simpl; auto with zarith.
+rewrite Zabs_eq; auto with zarith.
+rewrite inj_minus1; auto with zarith.
+rewrite inj_Zabs_nat; auto with zarith.
+rewrite Zabs_eq; auto with zarith.
+rewrite Zplus_comm; simpl; auto with zarith.
+apply inj_le_rev; rewrite inj_Zabs_nat; simpl; auto with zarith.
+rewrite Zabs_eq; auto with zarith.
+exists 0; auto with arith.
+Qed.
+
+End EGroup.
+
+Theorem gpow_gpow: forall (A : Set) (op : A -> A -> A) (a : A) (G : FGroup op),
+ In a (s G) -> forall n m, 0 <= n -> 0 <= m -> gpow a G (n * m ) = gpow (gpow a G n) G m.
+intros A op a G H n m; case n.
+simpl; intros _ H1; generalize H1.
+pattern m; apply natlike_ind; simpl; auto.
+intros x H2 Rec _; unfold Zsucc; rewrite gpow_add; simpl; auto with zarith.
+repeat rewrite G.(e_is_zero_r); auto with zarith.
+apply gpow_in; sauto.
+intros p1 _; case m; simpl; auto.
+assert(H1: In (iter_pos p1 A (op a) (e G)) (s G)).
+refine (gpow_in _ _ _ _ _ (Zpos p1)); auto.
+intros p2 _; pattern p2; apply Pind; simpl; auto.
+rewrite Pmult_1_r; rewrite G.(e_is_zero_r); try rewrite G.(e_is_zero_r); auto.
+intros p3 Rec; rewrite Pplus_one_succ_r; rewrite Pmult_plus_distr_l.
+rewrite Pmult_1_r.
+simpl; repeat rewrite iter_pos_plus; simpl.
+rewrite G.(e_is_zero_r); auto.
+rewrite gpow_op with (G:= G); try rewrite Rec; auto.
+apply sym_equal; apply gpow_op; auto.
+intros p Hp; contradict Hp; auto with zarith.
+Qed.
+
+Theorem gpow_e: forall (A : Set) (op : A -> A -> A) (G : FGroup op) n, 0 <= n -> gpow G.(e) G n = G.(e).
+intros A op G n; case n; simpl; auto with zarith.
+intros p _; elim p; simpl; auto; intros p1 Rec; repeat rewrite Rec; auto.
+Qed.
+
+Theorem gpow_pow: forall (A : Set) (op : A -> A -> A) (a : A) (G : FGroup op),
+ In a (s G) -> forall n, 0 <= n -> gpow a G (2 ^ n) = G.(e) -> forall m, n <= m -> gpow a G (2 ^ m) = G.(e).
+intros A op a G H n H1 H2 m Hm.
+replace m with (n + (m - n)); auto with zarith.
+rewrite Zpower_exp; auto with zarith.
+rewrite gpow_gpow; auto with zarith.
+rewrite H2; apply gpow_e.
+apply Zpower_ge_0; auto with zarith.
+Qed.
+
+Theorem gpow_mult: forall (A : Set) (op : A -> A -> A) (a b: A) (G : FGroup op)
+ (comm: forall a b, In a (s G) -> In b (s G) -> op a b = op b a),
+ In a (s G) -> In b (s G) -> forall n, 0 <= n -> gpow (op a b) G n = op (gpow a G n) (gpow b G n).
+intros A op a b G comm Ha Hb n; case n; simpl; auto.
+intros _; rewrite G.(e_is_zero_r); auto.
+2: intros p Hp; contradict Hp; auto with zarith.
+intros p _; pattern p; apply Pind; simpl; auto.
+repeat rewrite G.(e_is_zero_r); auto.
+intros p3 Rec; rewrite Pplus_one_succ_r.
+repeat rewrite iter_pos_plus; simpl.
+repeat rewrite (fun x y H z => gpow_op A op x G H (op y z)) ; auto.
+rewrite Rec.
+repeat rewrite G.(e_is_zero_r); auto.
+assert(H1: In (iter_pos p3 A (op a) (e G)) (s G)).
+refine (gpow_in _ _ _ _ _ (Zpos p3)); auto.
+assert(H2: In (iter_pos p3 A (op b) (e G)) (s G)).
+refine (gpow_in _ _ _ _ _ (Zpos p3)); auto.
+repeat rewrite <- G.(assoc); try eq_tac; auto.
+rewrite (fun x y => comm (iter_pos p3 A x y) b); auto.
+rewrite (G.(assoc) a); try apply comm; auto.
+Qed.
+
+Theorem Zdivide_mult_rel_prime: forall a b c : Z, (a | c) -> (b | c) -> rel_prime a b -> (a * b | c).
+intros a b c (q1, H1) (q2, H2) H3.
+assert (H4: (a | q2)).
+apply Gauss with (2 := H3).
+exists q1; rewrite <- H1; rewrite H2; auto with zarith.
+case H4; intros q3 H5; exists q3; rewrite H2; rewrite H5; auto with zarith.
+Qed.
+
+Theorem order_mult: forall (A : Set) (op : A -> A -> A) (A_dec: forall a b: A, {a = b} + {~ a = b}) (G : FGroup op)
+ (comm: forall a b, In a (s G) -> In b (s G) -> op a b = op b a) (a b: A),
+ In a (s G) -> In b (s G) -> rel_prime (e_order A_dec a G) (e_order A_dec b G) ->
+ e_order A_dec (op a b) G = e_order A_dec a G * e_order A_dec b G.
+intros A op A_dec G comm a b Ha Hb Hab.
+assert (Hoat: 0 < e_order A_dec a G); try apply e_order_pos.
+assert (Hobt: 0 < e_order A_dec b G); try apply e_order_pos.
+assert (Hoabt: 0 < e_order A_dec (op a b) G); try apply e_order_pos.
+assert (Hoa: 0 <= e_order A_dec a G); auto with zarith.
+assert (Hob: 0 <= e_order A_dec b G); auto with zarith.
+apply Zle_antisym; apply Zdivide_le; auto with zarith.
+apply Zmult_lt_O_compat; auto.
+apply e_order_divide_gpow; sauto; auto with zarith.
+rewrite gpow_mult; auto with zarith.
+rewrite gpow_gpow; auto with zarith.
+rewrite gpow_e_order_is_e; auto with zarith.
+rewrite gpow_e; auto.
+rewrite Zmult_comm.
+rewrite gpow_gpow; auto with zarith.
+rewrite gpow_e_order_is_e; auto with zarith.
+rewrite gpow_e; auto.
+apply Zdivide_mult_rel_prime; auto.
+apply Gauss with (2 := Hab).
+apply e_order_divide_gpow; auto with zarith.
+rewrite <- (gpow_e _ _ G (e_order A_dec b G)); auto.
+rewrite <- (gpow_e_order_is_e _ A_dec _ (op a b) G); auto with zarith.
+rewrite <- gpow_gpow; auto with zarith.
+rewrite (Zmult_comm (e_order A_dec (op a b) G)).
+rewrite gpow_mult; auto with zarith.
+rewrite gpow_gpow with (a := b); auto with zarith.
+rewrite gpow_e_order_is_e; auto with zarith.
+rewrite gpow_e; auto with zarith.
+rewrite G.(e_is_zero_r); auto with zarith.
+apply gpow_in; auto.
+apply Gauss with (2 := rel_prime_sym _ _ Hab).
+apply e_order_divide_gpow; auto with zarith.
+rewrite <- (gpow_e _ _ G (e_order A_dec a G)); auto.
+rewrite <- (gpow_e_order_is_e _ A_dec _ (op a b) G); auto with zarith.
+rewrite <- gpow_gpow; auto with zarith.
+rewrite (Zmult_comm (e_order A_dec (op a b) G)).
+rewrite gpow_mult; auto with zarith.
+rewrite gpow_gpow with (a := a); auto with zarith.
+rewrite gpow_e_order_is_e; auto with zarith.
+rewrite gpow_e; auto with zarith.
+rewrite G.(e_is_zero_l); auto with zarith.
+apply gpow_in; auto.
+Qed.
+
+Theorem fermat_gen: forall (A : Set) (A_dec: forall (a b: A), {a = b} + {a <>b}) (op : A -> A -> A) (a: A) (G : FGroup op),
+ In a G.(s) -> gpow a G (g_order G) = G.(e).
+intros A A_dec op a G H.
+assert (H1: (e_order A_dec a G | g_order G)).
+apply e_order_divide_g_order; auto.
+case H1; intros q; intros Hq; rewrite Hq.
+assert (Hq1: 0 <= q).
+apply Zmult_le_reg_r with (e_order A_dec a G); auto with zarith.
+apply Zlt_gt; apply e_order_pos.
+rewrite Zmult_0_l; rewrite <- Hq; apply Zlt_le_weak; apply g_order_pos.
+rewrite Zmult_comm; rewrite gpow_gpow; auto with zarith.
+rewrite gpow_e_order_is_e; auto with zarith.
+apply gpow_e; auto.
+apply Zlt_le_weak; apply e_order_pos.
+Qed.
+
+Theorem order_div: forall (A : Set) (A_dec: forall (a b: A), {a = b} + {a <>b}) (op : A -> A -> A) (a: A) (G : FGroup op) m,
+ 0 < m -> (forall p, prime p -> (p | m) -> gpow a G (m / p) <> G.(e)) ->
+ In a G.(s) -> gpow a G m = G.(e) -> e_order A_dec a G = m.
+intros A Adec op a G m Hm H H1 H2.
+assert (F1: 0 <= m); auto with zarith.
+case (e_order_divide_gpow A Adec op a G H1 m F1 H2); intros q Hq.
+assert (F2: 1 <= q).
+ case (Zle_or_lt 0 q); intros HH.
+ case (Zle_lt_or_eq _ _ HH); auto with zarith.
+ intros HH1; generalize Hm; rewrite Hq; rewrite <- HH1;
+ auto with zarith.
+ assert (F2: 0 <= (- q) * e_order Adec a G); auto with zarith.
+ apply Zmult_le_0_compat; auto with zarith.
+ apply Zlt_le_weak; apply e_order_pos.
+ generalize F2; rewrite Zopp_mult_distr_l_reverse;
+ rewrite <- Hq; auto with zarith.
+case (Zle_lt_or_eq _ _ F2); intros H3; subst; auto with zarith.
+case (prime_dec q); intros Hq.
+ case (H q); auto with zarith.
+ rewrite Zmult_comm; rewrite Z_div_mult; auto with zarith.
+ apply gpow_e_order_is_e; auto.
+case (Zdivide_div_prime_le_square _ H3 Hq); intros r (Hr1, (Hr2, Hr3)).
+case (H _ Hr1); auto.
+ apply Zdivide_trans with (1 := Hr2).
+ apply Zdivide_factor_r.
+case Hr2; intros q1 Hq1; subst.
+assert (F3: 0 < r).
+ generalize (prime_ge_2 _ Hr1); auto with zarith.
+rewrite <- Zmult_assoc; rewrite Zmult_comm; rewrite <- Zmult_assoc;
+ rewrite Zmult_comm; rewrite Z_div_mult; auto with zarith.
+rewrite gpow_gpow; auto with zarith.
+ rewrite gpow_e_order_is_e; try rewrite gpow_e; auto.
+ apply Zmult_le_reg_r with r; auto with zarith.
+ apply Zlt_le_weak; apply e_order_pos.
+apply Zmult_le_reg_r with r; auto with zarith.
+Qed.
diff --git a/coqprime/PrimalityTest/Euler.v b/coqprime/PrimalityTest/Euler.v
new file mode 100644
index 000000000..06d92ce57
--- /dev/null
+++ b/coqprime/PrimalityTest/Euler.v
@@ -0,0 +1,88 @@
+
+(*************************************************************)
+(* This file is distributed under the terms of the *)
+(* GNU Lesser General Public License Version 2.1 *)
+(*************************************************************)
+(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
+(*************************************************************)
+
+(************************************************************************
+
+ Definition of the Euler Totient function
+
+*************************************************************************)
+Require Import ZArith.
+Require Export Znumtheory.
+Require Import Tactic.
+Require Export ZSum.
+
+Open Scope Z_scope.
+
+Definition phi n := Zsum 1 (n - 1) (fun x => if rel_prime_dec x n then 1 else 0).
+
+Theorem phi_def_with_0:
+ forall n, 1< n -> phi n = Zsum 0 (n - 1) (fun x => if rel_prime_dec x n then 1 else 0).
+intros n H; rewrite Zsum_S_left; auto with zarith.
+case (rel_prime_dec 0 n); intros H2.
+contradict H2; apply not_rel_prime_0; auto.
+rewrite Zplus_0_l; auto.
+Qed.
+
+Theorem phi_pos: forall n, 1 < n -> 0 < phi n.
+intros n H; unfold phi.
+case (Zle_lt_or_eq 2 n); auto with zarith; intros H1; subst.
+rewrite Zsum_S_left; simpl; auto with zarith.
+case (rel_prime_dec 1 n); intros H2.
+apply Zlt_le_trans with (1 + 0); auto with zarith.
+apply Zplus_le_compat_l.
+pattern 0 at 1; replace 0 with ((1 + (n - 1) - 2) * 0); auto with zarith.
+rewrite <- Zsum_c; auto with zarith.
+apply Zsum_le; auto with zarith.
+intros x H3; case (rel_prime_dec x n); auto with zarith.
+case H2; apply rel_prime_1; auto with zarith.
+rewrite Zsum_nn.
+case (rel_prime_dec (2 - 1) 2); auto with zarith.
+intros H1; contradict H1; apply rel_prime_1; auto with zarith.
+Qed.
+
+Theorem phi_le_n_minus_1: forall n, 1 < n -> phi n <= n - 1.
+intros n H; replace (n-1) with ((1 + (n - 1) - 1) * 1); auto with zarith.
+rewrite <- Zsum_c; auto with zarith.
+unfold phi; apply Zsum_le; auto with zarith.
+intros x H1; case (rel_prime_dec x n); auto with zarith.
+Qed.
+
+Theorem prime_phi_n_minus_1: forall n, prime n -> phi n = n - 1.
+intros n H; replace (n-1) with ((1 + (n - 1) - 1) * 1); auto with zarith.
+assert (Hu: 1 <= n - 1).
+assert (2 <= n); auto with zarith.
+apply prime_ge_2; auto.
+rewrite <- Zsum_c; auto with zarith; unfold phi; apply Zsum_ext; auto.
+intros x (H2, H3); case H; clear H; intros H H1.
+generalize (H1 x); case (rel_prime_dec x n); auto with zarith.
+intros H6 H7; contradict H6; apply H7; split; auto with zarith.
+Qed.
+
+Theorem phi_n_minus_1_prime: forall n, 1 < n -> phi n = n - 1 -> prime n.
+intros n H H1; case (prime_dec n); auto; intros H2.
+assert (H3: phi n < n - 1); auto with zarith.
+replace (n-1) with ((1 + (n - 1) - 1) * 1); auto with zarith.
+assert (Hu: 1 <= n - 1); auto with zarith.
+rewrite <- Zsum_c; auto with zarith; unfold phi; apply Zsum_lt; auto.
+intros x _; case (rel_prime_dec x n); auto with zarith.
+case not_prime_divide with n; auto.
+intros x (H3, H4); exists x; repeat split; auto with zarith.
+case (rel_prime_dec x n); auto with zarith.
+intros H5; absurd (x = 1 \/ x = -1); auto with zarith.
+case (Zis_gcd_unique x n x 1); auto.
+apply Zis_gcd_intro; auto; exists 1; auto with zarith.
+contradict H3; rewrite H1; auto with zarith.
+Qed.
+
+Theorem phi_divide_prime: forall n, 1 < n -> (n - 1 | phi n) -> prime n.
+intros n H1 H2; apply phi_n_minus_1_prime; auto.
+apply Zle_antisym.
+apply phi_le_n_minus_1; auto.
+apply Zdivide_le; auto; auto with zarith.
+apply phi_pos; auto.
+Qed.
diff --git a/coqprime/PrimalityTest/FGroup.v b/coqprime/PrimalityTest/FGroup.v
new file mode 100644
index 000000000..a55710e7c
--- /dev/null
+++ b/coqprime/PrimalityTest/FGroup.v
@@ -0,0 +1,123 @@
+
+(*************************************************************)
+(* This file is distributed under the terms of the *)
+(* GNU Lesser General Public License Version 2.1 *)
+(*************************************************************)
+(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
+(*************************************************************)
+
+(**********************************************************************
+ FGroup.v
+
+ Defintion and properties of finite groups
+
+ Definition: FGroup
+ **********************************************************************)
+Require Import List.
+Require Import UList.
+Require Import Tactic.
+Require Import ZArith.
+
+Open Scope Z_scope.
+
+Set Implicit Arguments.
+
+(**************************************
+ A finite group is defined for an operation op
+ it has a support (s)
+ op operates inside the group (internal)
+ op is associative (assoc)
+ it has an element (e) that is neutral (e_is_zero_l e_is_zero_r)
+ it has an inverse operator (i)
+ the inverse operates inside the group (i_internal)
+ it gives an inverse (i_is_inverse_l is_is_inverse_r)
+ **************************************)
+
+Record FGroup (A: Set) (op: A -> A -> A): Set := mkGroup
+ {s : (list A);
+ unique_s: ulist s;
+ internal: forall a b, In a s -> In b s -> In (op a b) s;
+ assoc: forall a b c, In a s -> In b s -> In c s -> op a (op b c) = op (op a b) c;
+ e: A;
+ e_in_s: In e s;
+ e_is_zero_l: forall a, In a s -> op e a = a;
+ e_is_zero_r: forall a, In a s -> op a e = a;
+ i: A -> A;
+ i_internal: forall a, In a s -> In (i a) s;
+ i_is_inverse_l: forall a, (In a s) -> op (i a) a = e;
+ i_is_inverse_r: forall a, (In a s) -> op a (i a) = e
+}.
+
+(**************************************
+ The order of a group is the lengh of the support
+ **************************************)
+
+Definition g_order (A: Set) (op: A -> A -> A) (g: FGroup op) := Z_of_nat (length g.(s)).
+
+Unset Implicit Arguments.
+
+Hint Resolve unique_s internal e_in_s e_is_zero_l e_is_zero_r i_internal
+ i_is_inverse_l i_is_inverse_r assoc.
+
+
+Section FGroup.
+
+Variable A: Set.
+Variable op: A -> A -> A.
+
+(**************************************
+ Some properties of a finite group
+ **************************************)
+
+Theorem g_cancel_l: forall (g : FGroup op), forall a b c, In a g.(s) -> In b g.(s) -> In c g.(s) -> op a b = op a c -> b = c.
+intros g a b c H1 H2 H3 H4; apply trans_equal with (op g.(e) b); sauto.
+replace (g.(e)) with (op (g.(i) a) a); sauto.
+apply trans_equal with (op (i g a) (op a b)); sauto.
+apply sym_equal; apply assoc with g; auto.
+rewrite H4.
+apply trans_equal with (op (op (i g a) a) c); sauto.
+apply assoc with g; auto.
+replace (op (g.(i) a) a) with g.(e); sauto.
+Qed.
+
+Theorem g_cancel_r: forall (g : FGroup op), forall a b c, In a g.(s) -> In b g.(s) -> In c g.(s) -> op b a = op c a -> b = c.
+intros g a b c H1 H2 H3 H4; apply trans_equal with (op b g.(e)); sauto.
+replace (g.(e)) with (op a (g.(i) a)); sauto.
+apply trans_equal with (op (op b a) (i g a)); sauto.
+apply assoc with g; auto.
+rewrite H4.
+apply trans_equal with (op c (op a (i g a))); sauto.
+apply sym_equal; apply assoc with g; sauto.
+replace (op a (g.(i) a)) with g.(e); sauto.
+Qed.
+
+Theorem e_unique: forall (g : FGroup op), forall e1, In e1 g.(s) -> (forall a, In a g.(s) -> op e1 a = a) -> e1 = g.(e).
+intros g e1 He1 H2.
+apply trans_equal with (op e1 g.(e)); sauto.
+Qed.
+
+Theorem inv_op: forall (g: FGroup op) a b, In a g.(s) -> In b g.(s) -> g.(i) (op a b) = op (g.(i) b) (g.(i) a).
+intros g a1 b1 H1 H2; apply g_cancel_l with (g := g) (a := op a1 b1); sauto.
+repeat rewrite g.(assoc); sauto.
+apply trans_equal with g.(e); sauto.
+rewrite <- g.(assoc) with (a := a1); sauto.
+rewrite g.(i_is_inverse_r); sauto.
+rewrite g.(e_is_zero_r); sauto.
+Qed.
+
+Theorem i_e: forall (g: FGroup op), g.(i) g.(e) = g.(e).
+intro g; apply g_cancel_l with (g:= g) (a := g.(e)); sauto.
+apply trans_equal with g.(e); sauto.
+Qed.
+
+(**************************************
+ A group has at least one element
+ **************************************)
+
+Theorem g_order_pos: forall g: FGroup op, 0 < g_order g.
+intro g; generalize g.(e_in_s); unfold g_order; case g.(s); simpl; auto with zarith.
+Qed.
+
+
+
+End FGroup.
diff --git a/coqprime/PrimalityTest/IGroup.v b/coqprime/PrimalityTest/IGroup.v
new file mode 100644
index 000000000..11a73d414
--- /dev/null
+++ b/coqprime/PrimalityTest/IGroup.v
@@ -0,0 +1,253 @@
+
+(*************************************************************)
+(* This file is distributed under the terms of the *)
+(* GNU Lesser General Public License Version 2.1 *)
+(*************************************************************)
+(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
+(*************************************************************)
+
+(**********************************************************************
+ Igroup
+
+ Build the group of the inversible elements for the operation
+
+ Definition: ZpGroup
+ **********************************************************************)
+Require Import ZArith.
+Require Import Tactic.
+Require Import Wf_nat.
+Require Import UList.
+Require Import ListAux.
+Require Import FGroup.
+
+Open Scope Z_scope.
+
+Section IG.
+
+Variable A: Set.
+Variable op: A -> A -> A.
+Variable support: list A.
+Variable e: A.
+
+Hypothesis A_dec: forall a b: A, {a = b} + {a <> b}.
+Hypothesis support_ulist: ulist support.
+Hypothesis e_in_support: In e support.
+Hypothesis op_internal: forall a b, In a support -> In b support -> In (op a b) support.
+Hypothesis op_assoc: forall a b c, In a support -> In b support -> In c support -> op a (op b c) = op (op a b) c.
+Hypothesis e_is_zero_l: forall a, In a support -> op e a = a.
+Hypothesis e_is_zero_r: forall a, In a support -> op a e = a.
+
+(**************************************
+ is_inv_aux tests if there is an inverse of a for op in l
+ **************************************)
+
+Fixpoint is_inv_aux (l: list A) (a: A) {struct l}: bool :=
+ match l with nil => false | cons b l1 =>
+ if (A_dec (op a b) e) then if (A_dec (op b a) e) then true else is_inv_aux l1 a else is_inv_aux l1 a
+ end.
+
+Theorem is_inv_aux_false: forall b l, (forall a, (In a l) -> op b a <> e \/ op a b <> e) -> is_inv_aux l b = false.
+intros b l; elim l; simpl; auto.
+intros a l1 Rec H; case (A_dec (op a b) e); case (A_dec (op b a) e); auto.
+intros H1 H2; case (H a); auto; intros H3; case H3; auto.
+Qed.
+
+(**************************************
+ is_inv tests if there is an inverse in support
+ **************************************)
+Definition is_inv := is_inv_aux support.
+
+(**************************************
+ isupport_aux returns the sublist of inversible element of support
+ **************************************)
+
+Fixpoint isupport_aux (l: list A) : list A :=
+ match l with nil => nil | cons a l1 => if is_inv a then a::isupport_aux l1 else isupport_aux l1 end.
+
+(**************************************
+ Some properties of isupport_aux
+ **************************************)
+
+Theorem isupport_aux_is_inv_true: forall l a, In a (isupport_aux l) -> is_inv a = true.
+intros l a; elim l; simpl; auto.
+intros b l1 H; case_eq (is_inv b); intros H1; simpl; auto.
+intros [H2 | H2]; subst; auto.
+Qed.
+
+Theorem isupport_aux_is_in: forall l a, is_inv a = true -> In a l -> In a (isupport_aux l).
+intros l a; elim l; simpl; auto.
+intros b l1 Rec H [H1 | H1]; subst.
+rewrite H; auto with datatypes.
+case (is_inv b); auto with datatypes.
+Qed.
+
+
+Theorem isupport_aux_not_in:
+ forall b l, (forall a, (In a support) -> op b a <> e \/ op a b <> e) -> ~ In b (isupport_aux l).
+intros b l; elim l; simpl; simpl; auto.
+intros a l1 H; case_eq (is_inv a); intros H1; simpl; auto.
+intros H2 [H3 | H3]; subst.
+contradict H1.
+unfold is_inv; rewrite is_inv_aux_false; auto.
+case H; auto; apply isupport_aux_is_in; auto.
+Qed.
+
+Theorem isupport_aux_incl: forall l, incl (isupport_aux l) l.
+intros l; elim l; simpl; auto with datatypes.
+intros a l1 H1; case (is_inv a); auto with datatypes.
+Qed.
+
+Theorem isupport_aux_ulist: forall l, ulist l -> ulist (isupport_aux l).
+intros l; elim l; simpl; auto with datatypes.
+intros a l1 H1 H2; case_eq (is_inv a); intros H3; auto with datatypes.
+apply ulist_cons; auto with datatypes.
+intros H4; apply (ulist_app_inv _ (a::nil) l1 a); auto with datatypes.
+apply (isupport_aux_incl l1 a); auto.
+apply H1; apply ulist_app_inv_r with (a:: nil); auto.
+apply H1; apply ulist_app_inv_r with (a:: nil); auto.
+Qed.
+
+(**************************************
+ isupport is the sublist of inversible element of support
+ **************************************)
+
+Definition isupport := isupport_aux support.
+
+(**************************************
+ Some properties of isupport
+ **************************************)
+
+Theorem isupport_is_inv_true: forall a, In a isupport -> is_inv a = true.
+unfold isupport; intros a H; apply isupport_aux_is_inv_true with (1 := H).
+Qed.
+
+Theorem isupport_is_in: forall a, is_inv a = true -> In a support -> In a isupport.
+intros a H H1; unfold isupport; apply isupport_aux_is_in; auto.
+Qed.
+
+Theorem isupport_incl: incl isupport support.
+unfold isupport; apply isupport_aux_incl.
+Qed.
+
+Theorem isupport_ulist: ulist isupport.
+unfold isupport; apply isupport_aux_ulist.
+apply support_ulist.
+Qed.
+
+Theorem isupport_length: (length isupport <= length support)%nat.
+apply ulist_incl_length.
+apply isupport_ulist.
+apply isupport_incl.
+Qed.
+
+Theorem isupport_length_strict:
+ forall b, (In b support) -> (forall a, (In a support) -> op b a <> e \/ op a b <> e) ->
+ (length isupport < length support)%nat.
+intros b H H1; apply ulist_incl_length_strict.
+apply isupport_ulist.
+apply isupport_incl.
+intros H2; case (isupport_aux_not_in b support); auto.
+Qed.
+
+Fixpoint inv_aux (l: list A) (a: A) {struct l}: A :=
+ match l with nil => e | cons b l1 =>
+ if A_dec (op a b) e then if (A_dec (op b a) e) then b else inv_aux l1 a else inv_aux l1 a
+ end.
+
+Theorem inv_aux_prop_r: forall l a, is_inv_aux l a = true -> op a (inv_aux l a) = e.
+intros l a; elim l; simpl.
+intros; discriminate.
+intros b l1 H1; case (A_dec (op a b) e); case (A_dec (op b a) e); intros H3 H4; subst; auto.
+Qed.
+
+Theorem inv_aux_prop_l: forall l a, is_inv_aux l a = true -> op (inv_aux l a) a = e.
+intros l a; elim l; simpl.
+intros; discriminate.
+intros b l1 H1; case (A_dec (op a b) e); case (A_dec (op b a) e); intros H3 H4; subst; auto.
+Qed.
+
+Theorem inv_aux_inv: forall l a b, op a b = e -> op b a = e -> (In a l) -> is_inv_aux l b = true.
+intros l a b; elim l; simpl.
+intros _ _ H; case H.
+intros c l1 Rec H H0 H1; case H1; clear H1; intros H1; subst; rewrite H.
+case (A_dec (op b a) e); case (A_dec e e); auto.
+intros H1 H2; contradict H2; rewrite H0; auto.
+case (A_dec (op b c) e); case (A_dec (op c b) e); auto.
+Qed.
+
+Theorem inv_aux_in: forall l a, In (inv_aux l a) l \/ inv_aux l a = e.
+intros l a; elim l; simpl; auto.
+intros b l1; case (A_dec (op a b) e); case (A_dec (op b a) e); intros _ _ [H1 | H1]; auto.
+Qed.
+
+(**************************************
+ The inverse function
+ **************************************)
+
+Definition inv := inv_aux support.
+
+(**************************************
+ Some properties of inv
+ **************************************)
+
+Theorem inv_prop_r: forall a, In a isupport -> op a (inv a) = e.
+intros a H; unfold inv; apply inv_aux_prop_r with (l := support).
+change (is_inv a = true).
+apply isupport_is_inv_true; auto.
+Qed.
+
+Theorem inv_prop_l: forall a, In a isupport -> op (inv a) a = e.
+intros a H; unfold inv; apply inv_aux_prop_l with (l := support).
+change (is_inv a = true).
+apply isupport_is_inv_true; auto.
+Qed.
+
+Theorem is_inv_true: forall a b, op b a = e -> op a b = e -> (In a support) -> is_inv b = true.
+intros a b H H1 H2; unfold is_inv; apply inv_aux_inv with a; auto.
+Qed.
+
+Theorem is_inv_false: forall b, (forall a, (In a support) -> op b a <> e \/ op a b <> e) -> is_inv b = false.
+intros b H; unfold is_inv; apply is_inv_aux_false; auto.
+Qed.
+
+Theorem inv_internal: forall a, In a isupport -> In (inv a) isupport.
+intros a H; apply isupport_is_in.
+apply is_inv_true with a; auto.
+apply inv_prop_l; auto.
+apply inv_prop_r; auto.
+apply (isupport_incl a); auto.
+case (inv_aux_in support a); unfold inv; auto.
+intros H1; rewrite H1; apply e_in_support; auto with zarith.
+Qed.
+
+(**************************************
+ We are now ready to build our group
+ **************************************)
+
+Definition IGroup : (FGroup op).
+generalize (fun x=> (isupport_incl x)); intros Hx.
+apply mkGroup with (s := isupport) (e := e) (i := inv); auto.
+apply isupport_ulist.
+intros a b H H1.
+assert (Haii: In (inv a) isupport); try apply inv_internal; auto.
+assert (Hbii: In (inv b) isupport); try apply inv_internal; auto.
+apply isupport_is_in; auto.
+apply is_inv_true with (op (inv b) (inv a)); auto.
+rewrite op_assoc; auto.
+rewrite <- (op_assoc a); auto.
+rewrite inv_prop_r; auto.
+rewrite e_is_zero_r; auto.
+apply inv_prop_r; auto.
+rewrite <- (op_assoc (inv b)); auto.
+rewrite (op_assoc (inv a)); auto.
+rewrite inv_prop_l; auto.
+rewrite e_is_zero_l; auto.
+apply inv_prop_l; auto.
+apply isupport_is_in; auto.
+apply is_inv_true with e; auto.
+intros a H; apply inv_internal; auto.
+intros; apply inv_prop_l; auto.
+intros; apply inv_prop_r; auto.
+Defined.
+
+End IG.
diff --git a/coqprime/PrimalityTest/Lagrange.v b/coqprime/PrimalityTest/Lagrange.v
new file mode 100644
index 000000000..b35460bad
--- /dev/null
+++ b/coqprime/PrimalityTest/Lagrange.v
@@ -0,0 +1,179 @@
+
+(*************************************************************)
+(* This file is distributed under the terms of the *)
+(* GNU Lesser General Public License Version 2.1 *)
+(*************************************************************)
+(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
+(*************************************************************)
+
+(**********************************************************************
+ Lagrange.v
+
+ Proof of Lagrange theorem:
+ the oder of a subgroup divides the order of a group
+
+ Definition: lagrange
+ **********************************************************************)
+Require Import List.
+Require Import UList.
+Require Import ListAux.
+Require Import ZArith Znumtheory.
+Require Import NatAux.
+Require Import FGroup.
+
+Open Scope Z_scope.
+
+Section Lagrange.
+
+Variable A: Set.
+
+Variable A_dec: forall a b: A, {a = b} + {~ a = b}.
+
+Variable op: A -> A -> A.
+
+Variable G: (FGroup op).
+
+Variable H:(FGroup op).
+
+Hypothesis G_in_H: (incl G.(s) H.(s)).
+
+(**************************************
+ A group and a subgroup have the same neutral element
+ **************************************)
+
+Theorem same_e_for_H_and_G: H.(e) = G.(e).
+apply trans_equal with (op H.(e) H.(e)); sauto.
+apply trans_equal with (op H.(e) (op G.(e) (H.(i) G.(e)))); sauto.
+eq_tac; sauto.
+apply trans_equal with (op G.(e) (op G.(e) (H.(i) G.(e)))); sauto.
+repeat rewrite H.(assoc); sauto.
+eq_tac; sauto.
+apply trans_equal with G.(e); sauto.
+apply trans_equal with (op G.(e) H.(e)); sauto.
+eq_tac; sauto.
+Qed.
+
+(**************************************
+ The proof works like this.
+ If G = {e, g1, g2, g3, .., gn} and {e, h1, h2, h3, ..., hm}
+ we construct the list mkGH
+ {e, g1, g2, g3, ...., gn
+ hi*e, hi * g1, hi * g2, ..., hi * gn if hi does not appear before
+ ....
+ hk*e, hk * g1, hk * g2, ..., hk * gn if hk does not appear before
+ }
+ that contains all the element of H.
+ We show that this list does not contain double (ulist).
+ **************************************)
+
+Fixpoint mkList (base l: (list A)) { struct l} : (list A) :=
+ match l with
+ nil => nil
+ | cons a l1 => let r1 := mkList base l1 in
+ if (In_dec A_dec a r1) then r1 else
+ (map (op a) base) ++ r1
+ end.
+
+Definition mkGH := mkList G.(s) H.(s).
+
+Theorem mkGH_length: divide (length G.(s)) (length mkGH).
+unfold mkGH; elim H.(s); simpl.
+exists 0%nat; auto with arith.
+intros a l1 (c, H1); case (In_dec A_dec a (mkList G.(s) l1)); intros H2.
+exists c; auto.
+exists (1 + c)%nat; rewrite ListAux.length_app; rewrite ListAux.length_map; rewrite H1; ring.
+Qed.
+
+Theorem mkGH_incl: incl H.(s) mkGH.
+assert (H1: forall l, incl l H.(s) -> incl l (mkList G.(s) l)).
+intros l; elim l; simpl; auto with datatypes.
+intros a l1 H1 H2.
+case (In_dec A_dec a (mkList (s G) l1)); auto with datatypes.
+intros H3; assert (H4: incl l1 (mkList (s G) l1)).
+apply H1; auto with datatypes.
+intros b H4; apply H2; auto with datatypes.
+intros b; simpl; intros [H5 | H5]; subst; auto.
+intros _ b; simpl; intros [H3 | H3]; subst; auto.
+apply in_or_app; left.
+cut (In H.(e) G.(s)).
+elim (s G); simpl; auto.
+intros c l2 Hl2 [H3 | H3]; subst; sauto.
+assert (In b H.(s)); sauto.
+apply (H2 b); auto with datatypes.
+rewrite same_e_for_H_and_G; sauto.
+apply in_or_app; right.
+apply H1; auto with datatypes.
+apply incl_tran with (2:= H2); auto with datatypes.
+unfold mkGH; apply H1; auto with datatypes.
+Qed.
+
+Theorem incl_mkGH: incl mkGH H.(s).
+assert (H1: forall l, incl l H.(s) -> incl (mkList G.(s) l) H.(s)).
+intros l; elim l; simpl; auto with datatypes.
+intros a l1 H1 H2.
+case (In_dec A_dec a (mkList (s G) l1)); intros H3; auto with datatypes.
+apply H1; apply incl_tran with (2 := H2); auto with datatypes.
+apply incl_app.
+intros b H4.
+case ListAux.in_map_inv with (1:= H4); auto.
+intros c (Hc1, Hc2); subst; sauto.
+apply internal; auto with datatypes.
+apply H1; apply incl_tran with (2 := H2); auto with datatypes.
+unfold mkGH; apply H1; auto with datatypes.
+Qed.
+
+Theorem ulist_mkGH: ulist mkGH.
+assert (H1: forall l, incl l H.(s) -> ulist (mkList G.(s) l)).
+intros l; elim l; simpl; auto with datatypes.
+intros a l1 H1 H2.
+case (In_dec A_dec a (mkList (s G) l1)); intros H3; auto with datatypes.
+apply H1; apply incl_tran with (2 := H2); auto with datatypes.
+apply ulist_app; auto.
+apply ulist_map; sauto.
+intros x y H4 H5 H6; apply g_cancel_l with (g:= H) (a := a); sauto.
+apply H2; auto with datatypes.
+apply H1; apply incl_tran with (2 := H2); auto with datatypes.
+intros b H4 H5.
+case ListAux.in_map_inv with (1:= H4); auto.
+intros c (Hc, Hc1); subst.
+assert (H6: forall l a b, In b G.(s) -> incl l H.(s) -> In a (mkList G.(s) l) -> In (op a b) (mkList G.(s) l)).
+intros ll u v; elim ll; simpl; auto with datatypes.
+intros w ll1 T0 T1 T2.
+case (In_dec A_dec w (mkList (s G) ll1)); intros T3 T4; auto with datatypes.
+apply T0; auto; apply incl_tran with (2:= T2); auto with datatypes.
+case in_app_or with (1 := T4); intros T5; auto with datatypes.
+apply in_or_app; left.
+case ListAux.in_map_inv with (1:= T5); auto.
+intros z (Hz1, Hz2); subst.
+replace (op (op w z) v) with (op w (op z v)); sauto.
+apply in_map; sauto.
+apply assoc with H; auto with datatypes.
+apply in_or_app; right; auto with datatypes.
+apply T0; try apply incl_tran with (2 := T2); auto with datatypes.
+case H3; replace a with (op (op a c) (G.(i) c)); auto with datatypes.
+apply H6; sauto.
+apply incl_tran with (2 := H2); auto with datatypes.
+apply trans_equal with (op a (op c (G.(i) c))); sauto.
+apply sym_equal; apply assoc with H; auto with datatypes.
+replace (op c (G.(i) c)) with (G.(e)); sauto.
+rewrite <- same_e_for_H_and_G.
+assert (In a H.(s)); sauto; apply (H2 a); auto with datatypes.
+unfold mkGH; apply H1; auto with datatypes.
+Qed.
+
+(**************************************
+ Lagrange theorem
+ **************************************)
+
+Theorem lagrange: (g_order G | (g_order H)).
+unfold g_order.
+rewrite Permutation.permutation_length with (l := H.(s)) (m:= mkGH).
+case mkGH_length; intros x H1; exists (Z_of_nat x).
+rewrite H1; rewrite Zmult_comm; apply inj_mult.
+apply ulist_incl2_permutation; auto.
+apply ulist_mkGH.
+apply mkGH_incl.
+apply incl_mkGH.
+Qed.
+
+End Lagrange.
diff --git a/coqprime/PrimalityTest/LucasLehmer.v b/coqprime/PrimalityTest/LucasLehmer.v
new file mode 100644
index 000000000..c3c255036
--- /dev/null
+++ b/coqprime/PrimalityTest/LucasLehmer.v
@@ -0,0 +1,597 @@
+
+(*************************************************************)
+(* This file is distributed under the terms of the *)
+(* GNU Lesser General Public License Version 2.1 *)
+(*************************************************************)
+(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
+(*************************************************************)
+
+(**********************************************************************
+ LucasLehamer.v
+
+ Build the sequence for the primality test of Mersenne numbers
+
+ Definition: LucasLehmer
+ **********************************************************************)
+Require Import ZArith.
+Require Import ZCAux.
+Require Import Tactic.
+Require Import Wf_nat.
+Require Import NatAux.
+Require Import UList.
+Require Import ListAux.
+Require Import FGroup.
+Require Import EGroup.
+Require Import PGroup.
+Require Import IGroup.
+
+Open Scope Z_scope.
+
+(**************************************
+ The seeds of the serie
+ **************************************)
+
+Definition w := (2, 1).
+
+Definition v := (2, -1).
+
+Theorem w_plus_v: pplus w v = (4, 0).
+simpl; auto.
+Qed.
+
+Theorem w_mult_v : pmult w v = (1, 0).
+simpl; auto.
+Qed.
+
+(**************************************
+ Definition of the power function for pairs p^n
+ **************************************)
+
+Definition ppow p n := match n with Zpos q => iter_pos q _ (pmult p) (1, 0) | _ => (1, 0) end.
+
+(**************************************
+ Some properties of ppow
+ **************************************)
+
+Theorem ppow_0: forall n, ppow n 0 = (1, 0).
+simpl; auto.
+Qed.
+
+Theorem ppow_1: forall n, ppow (1, 0) n = (1, 0).
+intros n; case n; simpl; auto.
+intros p; apply iter_pos_invariant with (Inv := fun x => x = (1, 0)); auto.
+intros x H; rewrite H; auto.
+Qed.
+
+Theorem ppow_op: forall a b p, iter_pos p _ (pmult a) b = pmult (iter_pos p _ (pmult a) (1, 0)) b.
+intros a b p; generalize b; elim p; simpl; auto; clear b p.
+intros p Rec b.
+rewrite (Rec b).
+try rewrite (fun x y => Rec (pmult x y)); try rewrite (fun x y => Rec (iter_pos p _ x y)); auto.
+repeat rewrite pmult_assoc; auto.
+intros p Rec b.
+rewrite (Rec b); try rewrite (fun x y => Rec (pmult x y)); try rewrite (fun x y => Rec (iter_pos p _ x y)); auto.
+repeat rewrite pmult_assoc; auto.
+intros b; rewrite pmult_1_r; auto.
+Qed.
+
+Theorem ppow_add: forall n m p, 0 <= m -> 0 <= p -> ppow n (m + p) = pmult (ppow n m) (ppow n p).
+intros n m; case m; clear m.
+intros p _ _; rewrite ppow_0; rewrite pmult_1_l; auto.
+2: intros p m H; contradict H; auto with zarith.
+intros p1 m _; case m.
+intros _; rewrite Zplus_0_r; simpl; apply sym_equal; apply pmult_1_r.
+2: intros p2 H; contradict H; auto with zarith.
+intros p2 _; simpl.
+rewrite iter_pos_plus.
+rewrite ppow_op; auto.
+Qed.
+
+Theorem ppow_ppow: forall n m p, 0 <= n -> 0 <= m -> ppow p (n * m ) = ppow (ppow p n) m.
+intros n m; case n.
+intros p _ Hm; rewrite Zmult_0_l.
+rewrite ppow_0; apply sym_equal; apply ppow_1.
+2: intros p p1 H; contradict H; auto with zarith.
+intros p1 p _; case m; simpl; auto.
+intros p2 _; pattern p2; apply Pind; simpl; auto.
+rewrite Pmult_1_r; rewrite pmult_1_r; auto.
+intros p3 Rec; rewrite Pplus_one_succ_r; rewrite Pmult_plus_distr_l.
+rewrite Pmult_1_r.
+simpl; repeat rewrite iter_pos_plus; simpl.
+rewrite pmult_1_r.
+rewrite ppow_op; try rewrite Rec; auto.
+apply sym_equal; apply ppow_op; auto.
+Qed.
+
+
+Theorem ppow_mult: forall n m p, 0 <= n -> ppow (pmult m p) n = pmult (ppow m n) (ppow p n).
+intros n m p; case n; simpl; auto.
+intros p1 _; pattern p1; apply Pind; simpl; auto.
+repeat rewrite pmult_1_r; auto.
+intros p3 Rec; rewrite Pplus_one_succ_r.
+repeat rewrite iter_pos_plus; simpl.
+repeat rewrite (fun x y z => ppow_op x (pmult y z)) ; auto.
+rewrite Rec.
+repeat rewrite pmult_1_r; auto.
+repeat rewrite <- pmult_assoc; try eq_tac; auto.
+rewrite (fun x y => pmult_comm (iter_pos p3 _ x y) p); auto.
+rewrite (pmult_assoc m); try apply pmult_comm; auto.
+Qed.
+
+(**************************************
+ We can now define our series of pairs s
+ **************************************)
+
+Definition s n := pplus (ppow w (2 ^ n)) (ppow v (2 ^ n)).
+
+(**************************************
+ Some properties of s
+ **************************************)
+
+Theorem s0 : s 0 = (4, 0).
+simpl; auto.
+Qed.
+
+Theorem sn_aux: forall n, 0 <= n -> s (n+1) = (pplus (pmult (s n) (s n)) (-2, 0)).
+intros n Hn.
+assert (Hu: 0 <= 2 ^n); auto with zarith.
+set (y := (fst (s n) * fst (s n) - 2, 0)).
+unfold s; simpl; rewrite Zpower_exp; auto with zarith.
+rewrite Zpower_1_r; rewrite ppow_ppow; auto with zarith.
+repeat rewrite pplus_pmult_dist_r || rewrite pplus_pmult_dist_l.
+repeat rewrite <- pplus_assoc.
+eq_tac; auto.
+pattern 2 at 2; replace 2 with (1 + 1); auto with zarith.
+rewrite ppow_add; auto with zarith; simpl.
+rewrite pmult_1_r; auto.
+rewrite Zmult_comm; rewrite ppow_ppow; simpl; auto with zarith.
+repeat rewrite <- ppow_mult; auto with zarith.
+rewrite (pmult_comm v w); rewrite w_mult_v.
+rewrite ppow_1.
+repeat rewrite tpower_1.
+rewrite pplus_comm; repeat rewrite <- pplus_assoc;
+rewrite pplus_comm; repeat rewrite <- pplus_assoc.
+simpl; case (ppow (7, -4) (2 ^n)); simpl; intros z1 z2; eq_tac; auto with zarith.
+Qed.
+
+Theorem sn_snd: forall n, snd (s n) = 0.
+intros n; case n; simpl; auto.
+intros p; pattern p; apply Pind; auto.
+intros p1 H; rewrite Zpos_succ_morphism; unfold Zsucc.
+rewrite sn_aux; auto with zarith.
+generalize H; case (s (Zpos p1)); simpl.
+intros x y H1; rewrite H1; auto with zarith.
+Qed.
+
+Theorem sn: forall n, 0 <= n -> s (n+1) = (fst (s n) * fst (s n) -2, 0).
+intros n Hn; rewrite sn_aux; generalize (sn_snd n); case (s n); auto.
+intros x y H; simpl in H; rewrite H; simpl.
+eq_tac; ring.
+Qed.
+
+Theorem sn_w: forall n, 0 <= n -> ppow w (2 ^ (n + 1)) = pplus (pmult (s n) (ppow w (2 ^ n))) (- 1, 0).
+intros n H; unfold s; simpl; rewrite Zpower_exp; auto with zarith.
+assert (Hu: 0 <= 2 ^n); auto with zarith.
+rewrite Zpower_1_r; rewrite ppow_ppow; auto with zarith.
+repeat rewrite pplus_pmult_dist_r || rewrite pplus_pmult_dist_l.
+pattern 2 at 2; replace 2 with (1 + 1); auto with zarith.
+rewrite ppow_add; auto with zarith; simpl.
+rewrite pmult_1_r; auto.
+repeat rewrite <- ppow_mult; auto with zarith.
+rewrite (pmult_comm v w); rewrite w_mult_v.
+rewrite ppow_1; simpl.
+simpl; case (ppow (7, 4) (2 ^n)); simpl; intros z1 z2; eq_tac; auto with zarith.
+Qed.
+
+Theorem sn_w_next: forall n, 0 <= n -> ppow w (2 ^ (n + 1)) = pplus (pmult (s n) (ppow w (2 ^ n))) (- 1, 0).
+intros n H; unfold s; simpl; rewrite Zpower_exp; auto with zarith.
+assert (Hu: 0 <= 2 ^n); auto with zarith.
+rewrite Zpower_1_r; rewrite ppow_ppow; auto with zarith.
+repeat rewrite pplus_pmult_dist_r || rewrite pplus_pmult_dist_l.
+pattern 2 at 2; replace 2 with (1 + 1); auto with zarith.
+rewrite ppow_add; auto with zarith; simpl.
+rewrite pmult_1_r; auto.
+repeat rewrite <- ppow_mult; auto with zarith.
+rewrite (pmult_comm v w); rewrite w_mult_v.
+rewrite ppow_1; simpl.
+simpl; case (ppow (7, 4) (2 ^n)); simpl; intros z1 z2; eq_tac; auto with zarith.
+Qed.
+
+Section Lucas.
+
+Variable p: Z.
+
+(**************************************
+ Definition of the mersenne number
+ **************************************)
+
+Definition Mp := 2^p -1.
+
+Theorem mersenne_pos: 1 < p -> 1 < Mp.
+intros H; unfold Mp; assert (2 < 2 ^p); auto with zarith.
+apply Zlt_le_trans with (2^2); auto with zarith.
+refine (refl_equal _).
+apply Zpower_le_monotone; auto with zarith.
+Qed.
+
+Hypothesis p_pos2: 2 < p.
+
+(**************************************
+ We suppose that the mersenne number divides s
+ **************************************)
+
+Hypothesis Mp_divide_sn: (Mp | fst (s (p - 2))).
+
+Variable q: Z.
+
+(**************************************
+ We take a divisor of Mp and shows that Mp <= q^2, hence Mp is prime
+ **************************************)
+
+Hypothesis q_divide_Mp: (q | Mp).
+
+Hypothesis q_pos2: 2 < q.
+
+Theorem q_pos: 1 < q.
+apply Zlt_trans with (2 := q_pos2); auto with zarith.
+Qed.
+
+(**************************************
+ The definition of the groups of inversible pairs
+ **************************************)
+
+Definition pgroup := PGroup q q_pos.
+
+Theorem w_in_pgroup: (In w pgroup.(FGroup.s)).
+generalize q_pos; intros HM.
+generalize q_pos2; intros HM2.
+assert (H0: 0 < q); auto with zarith.
+simpl; apply isupport_is_in; auto.
+assert (zpmult q w (2, q - 1) = (1, 0)).
+unfold zpmult, w, pmult, base; repeat (rewrite Zmult_1_r || rewrite Zmult_1_l).
+eq_tac.
+apply trans_equal with ((3 * q + 1) mod q).
+eq_tac; auto with zarith.
+rewrite Zplus_mod; auto.
+rewrite Zmult_mod; auto.
+rewrite Z_mod_same; auto with zarith.
+rewrite Zmult_0_r; repeat rewrite Zmod_small; auto with zarith.
+apply trans_equal with (2 * q mod q).
+eq_tac; auto with zarith.
+apply Zdivide_mod; auto with zarith; exists 2; auto with zarith.
+apply is_inv_true with (2, q - 1); auto.
+apply mL_in; auto with zarith.
+intros; apply zpmult_1_l; auto with zarith.
+intros; apply zpmult_1_r; auto with zarith.
+rewrite zpmult_comm; auto.
+apply mL_in; auto with zarith.
+unfold w; apply mL_in; auto with zarith.
+Qed.
+
+Theorem e_order_divide_order: (e_order P_dec w pgroup | g_order pgroup).
+apply e_order_divide_g_order.
+apply w_in_pgroup.
+Qed.
+
+Theorem order_lt: g_order pgroup < q * q.
+unfold g_order, pgroup, PGroup; simpl.
+rewrite <- (Zabs_eq (q * q)); auto with zarith.
+rewrite <- (inj_Zabs_nat (q * q)); auto with zarith.
+rewrite <- mL_length; auto with zarith.
+apply inj_lt; apply isupport_length_strict with (0, 0).
+apply mL_ulist.
+apply mL_in; auto with zarith.
+intros a _; left; rewrite zpmult_0_l; auto with zarith.
+intros; discriminate.
+Qed.
+
+(**************************************
+ The power function zpow: a^n
+ **************************************)
+
+Definition zpow a := gpow a pgroup.
+
+(**************************************
+ Some properties of zpow
+ **************************************)
+
+Theorem zpow_def:
+ forall a b, In a pgroup.(FGroup.s) -> 0 <= b ->
+ zpow a b = ((fst (ppow a b)) mod q, (snd (ppow a b)) mod q).
+generalize q_pos; intros HM.
+generalize q_pos2; intros HM2.
+assert (H0: 0 < q); auto with zarith.
+intros a b Ha Hb; generalize Hb; pattern b; apply natlike_ind; auto.
+intros _; repeat rewrite Zmod_small; auto with zarith.
+rewrite ppow_0; simpl; auto with zarith.
+unfold zpow; intros n1 H Rec _; unfold Zsucc.
+rewrite gpow_add; auto with zarith.
+rewrite ppow_add; simpl; try rewrite pmult_1_r; auto with zarith.
+rewrite Rec; unfold zpmult; auto with zarith.
+case (ppow a n1); case a; unfold pmult, fst, snd.
+intros x y z t.
+repeat (rewrite Zmult_1_r || rewrite Zmult_0_r || rewrite Zplus_0_r || rewrite Zplus_0_l); eq_tac.
+repeat rewrite (fun u v => Zplus_mod (u * v)); auto.
+eq_tac; try eq_tac; auto.
+repeat rewrite (Zmult_mod z); auto with zarith.
+repeat rewrite (fun u v => Zmult_mod (u * v)); auto.
+eq_tac; try eq_tac; auto with zarith.
+repeat rewrite (Zmult_mod base); auto with zarith.
+eq_tac; try eq_tac; auto with zarith.
+apply Zmod_mod; auto.
+apply Zmod_mod; auto.
+repeat rewrite (fun u v => Zplus_mod (u * v)); auto.
+eq_tac; try eq_tac; auto.
+repeat rewrite (Zmult_mod z); auto with zarith.
+repeat rewrite (Zmult_mod t); auto with zarith.
+Qed.
+
+Theorem zpow_w_n_minus_1: zpow w (2 ^ (p - 1)) = (-1 mod q, 0).
+generalize q_pos; intros HM.
+generalize q_pos2; intros HM2.
+assert (H0: 0 < q); auto with zarith.
+rewrite zpow_def.
+replace (p - 1) with ((p - 2) + 1); auto with zarith.
+rewrite sn_w; auto with zarith.
+generalize Mp_divide_sn (sn_snd (p - 2)); case (s (p -2)); case (ppow w (2 ^ (p -2))).
+unfold fst, snd; intros x y z t H1 H2; unfold pmult, pplus; subst.
+repeat (rewrite Zmult_0_l || rewrite Zmult_0_r || rewrite Zplus_0_l || rewrite Zplus_0_r).
+assert (H2: z mod q = 0).
+case H1; intros q1 Hq1; rewrite Hq1.
+case q_divide_Mp; intros q2 Hq2; rewrite Hq2.
+rewrite Zmult_mod; auto.
+rewrite (Zmult_mod q2); auto.
+rewrite Z_mod_same; auto with zarith.
+repeat (rewrite Zmult_0_r; rewrite (Zmod_small 0)); auto with zarith.
+assert (H3: forall x, (z * x) mod q = 0).
+intros y1; rewrite Zmult_mod; try rewrite H2; auto.
+assert (H4: forall x y, (z * x + y) mod q = y mod q).
+intros x1 y1; rewrite Zplus_mod; try rewrite H3; auto.
+rewrite Zplus_0_l; apply Zmod_mod; auto.
+eq_tac; auto.
+apply w_in_pgroup.
+apply Zlt_le_weak; apply Zpower_gt_0; auto with zarith.
+Qed.
+
+Theorem zpow_w_n: zpow w (2 ^ p) = (1, 0).
+generalize q_pos; intros HM.
+generalize q_pos2; intros HM2.
+assert (H0: 0 < q); auto with zarith.
+replace p with ((p - 1) + 1); auto with zarith.
+rewrite Zpower_exp; try rewrite Zpower_exp_1; auto with zarith.
+unfold zpow; rewrite gpow_gpow; auto with zarith.
+generalize zpow_w_n_minus_1; unfold zpow; intros H1; rewrite H1; clear H1.
+simpl; unfold zpmult, pmult.
+repeat (rewrite Zmult_0_l || rewrite Zmult_0_r || rewrite Zplus_0_l ||
+ rewrite Zplus_0_r || rewrite Zmult_1_r).
+eq_tac; auto.
+pattern (-1 mod q) at 1; rewrite <- (Zmod_mod (-1) q); auto with zarith.
+repeat rewrite <- Zmult_mod; auto.
+rewrite Zmod_small; auto with zarith.
+apply w_in_pgroup.
+Qed.
+
+(**************************************
+ As e = (1, 0), the previous equation implies that the order of the group divide 2^p
+ **************************************)
+
+Theorem e_order_divide_pow: (e_order P_dec w pgroup | 2 ^ p).
+generalize q_pos; intros HM.
+generalize q_pos2; intros HM2.
+assert (H0: 0 < q); auto with zarith.
+apply e_order_divide_gpow.
+apply w_in_pgroup.
+apply Zlt_le_weak; apply Zpower_gt_0; auto with zarith.
+exact zpow_w_n.
+Qed.
+
+(**************************************
+ So it is less than equal
+ **************************************)
+
+Theorem e_order_le_pow : e_order P_dec w pgroup <= 2 ^ p.
+apply Zdivide_le.
+apply Zlt_le_weak; apply e_order_pos.
+apply Zpower_gt_0; auto with zarith.
+apply e_order_divide_pow.
+Qed.
+
+(**************************************
+ So order(w) must be 2^q
+ **************************************)
+
+Theorem e_order_eq_pow: exists q, (e_order P_dec w pgroup) = 2 ^ q.
+case (Zdivide_power_2 (e_order P_dec w pgroup) 2 p); auto with zarith.
+apply Zlt_le_weak; apply e_order_pos.
+apply prime_2.
+apply e_order_divide_pow; auto.
+intros x H; exists x; auto with zarith.
+Qed.
+
+(**************************************
+ Buth this q can only be p otherwise it would contradict w^2^(p -1) = (-1, 0)
+ **************************************)
+
+Theorem e_order_eq_p: e_order P_dec w pgroup = 2 ^ p.
+case (Zdivide_power_2 (e_order P_dec w pgroup) 2 p); auto with zarith.
+apply Zlt_le_weak; apply e_order_pos.
+apply prime_2.
+apply e_order_divide_pow; auto.
+intros p1 Hp1.
+case (Zle_lt_or_eq p1 p); try (intro H1; subst; auto; fail).
+case (Zle_or_lt p1 p); auto; intros H1.
+absurd (2 ^ p1 <= 2 ^ p); auto with zarith.
+apply Zlt_not_le; apply Zpower_lt_monotone; auto with zarith.
+apply Zdivide_le.
+apply Zlt_le_weak; apply Zpower_gt_0; auto with zarith.
+apply Zpower_gt_0; auto with zarith.
+rewrite <- Hp1; apply e_order_divide_pow.
+intros H1.
+assert (Hu: 0 <= p1).
+generalize Hp1; case p1; simpl; auto with zarith.
+intros p2 Hu; absurd (0 < e_order P_dec w pgroup).
+rewrite Hu; auto with zarith.
+apply e_order_pos.
+absurd (zpow w (2 ^ (p - 1)) = (1, 0)).
+rewrite zpow_w_n_minus_1.
+intros H2; injection H2; clear H2; intros H2.
+assert (H0: 0 < q); auto with zarith.
+absurd (0 mod q = 0).
+pattern 0 at 1; replace 0 with (-1 + 1); auto with zarith.
+rewrite Zplus_mod; auto with zarith.
+rewrite H2; rewrite (Zmod_small 1); auto with zarith.
+rewrite Zmod_small; auto with zarith.
+rewrite Zmod_small; auto with zarith.
+unfold zpow; apply (gpow_pow _ _ w pgroup) with p1; auto with zarith.
+apply w_in_pgroup.
+rewrite <- Hp1.
+apply (gpow_e_order_is_e _ P_dec _ w pgroup).
+apply w_in_pgroup.
+Qed.
+
+(**************************************
+ We have then the expected conclusion
+ **************************************)
+
+Theorem q_more_than_square: Mp < q * q.
+unfold Mp.
+assert (2 ^ p <= q * q); auto with zarith.
+rewrite <- e_order_eq_p.
+apply Zle_trans with (g_order pgroup).
+apply Zdivide_le; auto with zarith.
+apply Zlt_le_weak; apply e_order_pos; auto with zarith.
+2: apply e_order_divide_order.
+2: apply Zlt_le_weak; apply order_lt.
+apply Zlt_le_trans with 2; auto with zarith.
+replace 2 with (Z_of_nat (length ((1, 0)::w::nil))); auto.
+unfold g_order; apply inj_le.
+apply ulist_incl_length.
+apply ulist_cons; simpl; auto.
+unfold w; intros [H2 | H2]; try (case H2; fail); discriminate.
+intro a; simpl; intros [H1 | [H1 | H1]]; subst.
+assert (In (1, 0) (mL q)).
+apply mL_in; auto with zarith.
+apply isupport_is_in; auto.
+apply is_inv_true with (1, 0); simpl; auto.
+intros; apply zpmult_1_l; auto with zarith.
+intros; apply zpmult_1_r; auto with zarith.
+rewrite zpmult_1_r; auto with zarith.
+rewrite zpmult_1_r; auto with zarith.
+exact w_in_pgroup.
+case H1.
+Qed.
+
+End Lucas.
+
+(**************************************
+ We build the sequence in Z
+ **************************************)
+
+Definition SS p :=
+ let n := Mp p in
+ match p - 2 with
+ Zpos p1 => iter_pos p1 _ (fun x => Zmodd (Zsquare x - 2) n) (Zmodd 4 n)
+ | _ => (Zmodd 4 n)
+ end.
+
+Theorem SS_aux_correct:
+ forall p z1 z2 n, 0 <= n -> 0 < z1 -> z2 = fst (s n) mod z1 ->
+ iter_pos p _ (fun x => Zmodd (Zsquare x - 2) z1) z2 = fst (s (n + Zpos p)) mod z1.
+intros p; pattern p; apply Pind.
+simpl.
+intros z1 z2 n Hn H H1; rewrite sn; auto; rewrite H1; rewrite Zmodd_correct; rewrite Zsquare_correct; simpl.
+unfold Zminus; rewrite Zplus_mod; auto.
+rewrite (Zplus_mod (fst (s n) * fst (s n))); auto with zarith.
+eq_tac; auto.
+eq_tac; auto.
+apply sym_equal; apply Zmult_mod; auto.
+intros n Rec z1 z2 n1 Hn1 H1 H2.
+rewrite Pplus_one_succ_l; rewrite iter_pos_plus.
+rewrite Rec with (n0 := n1); auto.
+replace (n1 + Zpos (1 + n)) with ((n1 + Zpos n) + 1); auto with zarith.
+rewrite sn; simpl; try rewrite Zmodd_correct; try rewrite Zsquare_correct; simpl; auto with zarith.
+unfold Zminus; rewrite Zplus_mod; auto.
+unfold Zmodd.
+rewrite (Zplus_mod (fst (s (n1 + Zpos n)) * fst (s (n1 + Zpos n)))); auto with zarith.
+eq_tac; auto.
+eq_tac; auto.
+apply sym_equal; apply Zmult_mod; auto.
+rewrite Zpos_plus_distr; auto with zarith.
+Qed.
+
+Theorem SS_prop: forall n, 1 < n -> SS n = fst(s (n -2)) mod (Mp n).
+intros n Hn; unfold SS.
+cut (0 <= n - 2); auto with zarith.
+case (n - 2).
+intros _; rewrite Zmodd_correct; rewrite s0; auto.
+intros p1 H2; rewrite SS_aux_correct with (n := 0); auto with zarith.
+apply Zle_lt_trans with 1; try apply mersenne_pos; auto with zarith.
+rewrite Zmodd_correct; rewrite s0; auto.
+intros p1 H2; case H2; auto.
+Qed.
+
+Theorem SS_prop_cor: forall p, 1 < p -> SS p = 0 -> (Mp p | fst(s (p -2))).
+intros p H H1.
+apply Zmod_divide.
+generalize (mersenne_pos _ H); auto with zarith.
+apply trans_equal with (2:= H1); apply sym_equal; apply SS_prop; auto.
+Qed.
+
+Theorem LucasLehmer: forall p, 2 < p -> SS p = 0 -> prime (Mp p).
+intros p H H1; case (prime_dec (Mp p)); auto; intros H2.
+case Zdivide_div_prime_le_square with (2 := H2).
+apply mersenne_pos; apply Zlt_trans with 2; auto with zarith.
+intros q (H3, (H4, H5)).
+contradict H5; apply Zlt_not_le.
+apply q_more_than_square; auto.
+apply SS_prop_cor; auto.
+apply Zlt_trans with 2; auto with zarith.
+case (Zle_lt_or_eq 2 q); auto.
+apply prime_ge_2; auto.
+intros H5; subst.
+absurd (2 <= 1); auto with arith.
+apply Zdivide_le; auto with zarith.
+case H4; intros x Hx.
+exists (2 ^ (p -1) - x).
+rewrite Zmult_minus_distr_r; rewrite <- Hx; unfold Mp.
+pattern 2 at 2; rewrite <- Zpower_1_r; rewrite <- Zpower_exp; auto with zarith.
+replace (p - 1 + 1) with p; auto with zarith.
+Qed.
+
+(**************************************
+ The test
+ **************************************)
+
+Definition lucas_test n :=
+ if Z_lt_dec 2 n then if Z_eq_dec (SS n) 0 then true else false else false.
+
+Theorem LucasTest: forall n, lucas_test n = true -> prime (Mp n).
+intros n; unfold lucas_test; case (Z_lt_dec 2 n); intros H1; try (intros; discriminate).
+case (Z_eq_dec (SS n) 0); intros H2; try (intros; discriminate).
+intros _; apply LucasLehmer; auto.
+Qed.
+
+Theorem prime7: prime 7.
+exact (LucasTest 3 (refl_equal _)).
+Qed.
+
+Theorem prime31: prime 31.
+exact (LucasTest 5 (refl_equal _)).
+Qed.
+
+Theorem prime127: prime 127.
+exact (LucasTest 7 (refl_equal _)).
+Qed.
+
+Theorem prime8191: prime 8191.
+exact (LucasTest 13 (refl_equal _)).
+Qed.
+
+Theorem prime131071: prime 131071.
+exact (LucasTest 17 (refl_equal _)).
+Qed.
+
+Theorem prime524287: prime 524287.
+exact (LucasTest 19 (refl_equal _)).
+Qed.
+
diff --git a/coqprime/PrimalityTest/Makefile.bak b/coqprime/PrimalityTest/Makefile.bak
new file mode 100644
index 000000000..fe49dbf29
--- /dev/null
+++ b/coqprime/PrimalityTest/Makefile.bak
@@ -0,0 +1,203 @@
+##############################################################################
+## The Calculus of Inductive Constructions ##
+## ##
+## Projet Coq ##
+## ##
+## INRIA ENS-CNRS ##
+## Rocquencourt Lyon ##
+## ##
+## Coq V7 ##
+## ##
+## ##
+##############################################################################
+
+# WARNING
+#
+# This Makefile has been automagically generated by coq_makefile
+# Edit at your own risks !
+#
+# END OF WARNING
+
+#
+# This Makefile was generated by the command line :
+# coq_makefile -f Make -o Makefile
+#
+
+##########################
+# #
+# Variables definitions. #
+# #
+##########################
+
+CAMLP4LIB=`camlp4 -where`
+COQSRC=-I $(COQTOP)/kernel -I $(COQTOP)/lib \
+ -I $(COQTOP)/library -I $(COQTOP)/parsing \
+ -I $(COQTOP)/pretyping -I $(COQTOP)/interp \
+ -I $(COQTOP)/proofs -I $(COQTOP)/syntax -I $(COQTOP)/tactics \
+ -I $(COQTOP)/toplevel -I $(COQTOP)/contrib/correctness \
+ -I $(COQTOP)/contrib/extraction -I $(COQTOP)/contrib/field \
+ -I $(COQTOP)/contrib/fourier -I $(COQTOP)/contrib/graphs \
+ -I $(COQTOP)/contrib/interface -I $(COQTOP)/contrib/jprover \
+ -I $(COQTOP)/contrib/omega -I $(COQTOP)/contrib/romega \
+ -I $(COQTOP)/contrib/ring -I $(COQTOP)/contrib/xml \
+ -I $(CAMLP4LIB)
+ZFLAGS=$(OCAMLLIBS) $(COQSRC)
+OPT=
+COQFLAGS=-q $(OPT) $(COQLIBS) $(OTHERFLAGS) $(COQ_XML)
+COQC=$(COQBIN)coqc
+GALLINA=gallina
+COQDOC=coqdoc
+CAMLC=ocamlc -c
+CAMLOPTC=ocamlopt -c
+CAMLLINK=ocamlc
+CAMLOPTLINK=ocamlopt
+COQDEP=$(COQBIN)coqdep -c
+GRAMMARS=grammar.cma
+CAMLP4EXTEND=pa_extend.cmo pa_ifdef.cmo q_MLast.cmo
+PP=-pp "camlp4o -I . -I $(COQTOP)/parsing $(CAMLP4EXTEND) $(GRAMMARS) -impl"
+
+#########################
+# #
+# Libraries definition. #
+# #
+#########################
+
+OCAMLLIBS=-I .\
+ -I ../Tactic\
+ -I ../N\
+ -I ../Z\
+ -I ../List
+COQLIBS=-I .\
+ -I ../Tactic\
+ -I ../N\
+ -I ../Z\
+ -I ../List
+
+###################################
+# #
+# Definition of the "all" target. #
+# #
+###################################
+
+VFILES=Cyclic.v\
+ EGroup.v\
+ Euler.v\
+ FGroup.v\
+ IGroup.v\
+ Lagrange.v\
+ LucasLehmer.v\
+ Pepin.v\
+ PGroup.v\
+ PocklingtonCertificat.v\
+ PocklingtonRefl.v\
+ Pocklington.v\
+ Proth.v\
+ Root.v\
+ Zp.v
+VOFILES=$(VFILES:.v=.vo)
+VIFILES=$(VFILES:.v=.vi)
+GFILES=$(VFILES:.v=.g)
+HTMLFILES=$(VFILES:.v=.html)
+GHTMLFILES=$(VFILES:.v=.g.html)
+
+all: Cyclic.vo\
+ EGroup.vo\
+ Euler.vo\
+ FGroup.vo\
+ IGroup.vo\
+ Lagrange.vo\
+ LucasLehmer.vo\
+ Pepin.vo\
+ PGroup.vo\
+ PocklingtonCertificat.vo\
+ PocklingtonRefl.vo\
+ Pocklington.vo\
+ Proth.vo\
+ Root.vo\
+ Zp.vo
+
+spec: $(VIFILES)
+
+gallina: $(GFILES)
+
+html: $(HTMLFILES)
+
+gallinahtml: $(GHTMLFILES)
+
+all.ps: $(VFILES)
+ $(COQDOC) -ps -o $@ `$(COQDEP) -sort -suffix .v $(VFILES)`
+
+all-gal.ps: $(VFILES)
+ $(COQDOC) -ps -g -o $@ `$(COQDEP) -sort -suffix .v $(VFILES)`
+
+
+
+####################
+# #
+# Special targets. #
+# #
+####################
+
+.PHONY: all opt byte archclean clean install depend html
+
+.SUFFIXES: .v .vo .vi .g .html .tex .g.tex .g.html
+
+.v.vo:
+ $(COQC) $(COQDEBUG) $(COQFLAGS) $*
+
+.v.vi:
+ $(COQC) -i $(COQDEBUG) $(COQFLAGS) $*
+
+.v.g:
+ $(GALLINA) $<
+
+.v.tex:
+ $(COQDOC) -latex $< -o $@
+
+.v.html:
+ $(COQDOC) -html $< -o $@
+
+.v.g.tex:
+ $(COQDOC) -latex -g $< -o $@
+
+.v.g.html:
+ $(COQDOC) -html -g $< -o $@
+
+byte:
+ $(MAKE) all "OPT="
+
+opt:
+ $(MAKE) all "OPT=-opt"
+
+include .depend
+
+.depend depend:
+ rm -f .depend
+ $(COQDEP) -i $(COQLIBS) $(VFILES) *.ml *.mli >.depend
+ $(COQDEP) $(COQLIBS) -suffix .html $(VFILES) >>.depend
+
+install:
+ mkdir -p `$(COQC) -where`/user-contrib
+ cp -f $(VOFILES) `$(COQC) -where`/user-contrib
+
+Makefile: Make
+ mv -f Makefile Makefile.bak
+ $(COQBIN)coq_makefile -f Make -o Makefile
+
+
+clean:
+ rm -f *.cmo *.cmi *.cmx *.o $(VOFILES) $(VIFILES) $(GFILES) *~
+ rm -f all.ps all-gal.ps $(HTMLFILES) $(GHTMLFILES)
+
+archclean:
+ rm -f *.cmx *.o
+
+html:
+
+# WARNING
+#
+# This Makefile has been automagically generated by coq_makefile
+# Edit at your own risks !
+#
+# END OF WARNING
+
diff --git a/coqprime/PrimalityTest/Note.pdf b/coqprime/PrimalityTest/Note.pdf
new file mode 100644
index 000000000..239a38772
--- /dev/null
+++ b/coqprime/PrimalityTest/Note.pdf
Binary files differ
diff --git a/coqprime/PrimalityTest/PGroup.v b/coqprime/PrimalityTest/PGroup.v
new file mode 100644
index 000000000..e9c1b2f47
--- /dev/null
+++ b/coqprime/PrimalityTest/PGroup.v
@@ -0,0 +1,347 @@
+
+(*************************************************************)
+(* This file is distributed under the terms of the *)
+(* GNU Lesser General Public License Version 2.1 *)
+(*************************************************************)
+(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
+(*************************************************************)
+
+(**********************************************************************
+ PGroup.v
+
+ Build the group of pairs modulo needed for the theorem of
+ lucas lehmer
+
+ Definition: PGroup
+ **********************************************************************)
+Require Import ZArith.
+Require Import Znumtheory.
+Require Import Tactic.
+Require Import Wf_nat.
+Require Import ListAux.
+Require Import UList.
+Require Import FGroup.
+Require Import EGroup.
+Require Import IGroup.
+
+Open Scope Z_scope.
+
+Definition base := 3.
+
+
+(**************************************
+ Equality is decidable on pairs
+ **************************************)
+
+Definition P_dec: forall p q: Z * Z, {p = q} + {p <> q}.
+intros p1 q1; case p1; case q1; intros z t x y; case (Z_eq_dec x z); intros H1.
+case (Z_eq_dec y t); intros H2.
+left; eq_tac; auto.
+right; contradict H2; injection H2; auto.
+right; contradict H1; injection H1; auto.
+Defined.
+
+
+(**************************************
+ Addition of two pairs
+ **************************************)
+
+Definition pplus (p q: Z * Z) := let (x ,y) := p in let (z,t) := q in (x + z, y + t).
+
+(**************************************
+ Properties of addition
+ **************************************)
+
+Theorem pplus_assoc: forall p q r, (pplus p (pplus q r)) = (pplus (pplus p q) r).
+intros p q r; case p; case q; case r; intros r1 r2 q1 q2 p1 p2; unfold pplus.
+eq_tac; ring.
+Qed.
+
+Theorem pplus_comm: forall p q, (pplus p q) = (pplus q p).
+intros p q; case p; case q; intros q1 q2 p1 p2; unfold pplus.
+eq_tac; ring.
+Qed.
+
+(**************************************
+ Multiplication of two pairs
+ **************************************)
+
+Definition pmult (p q: Z * Z) := let (x ,y) := p in let (z,t) := q in (x * z + base * y * t, x * t + y * z).
+
+(**************************************
+ Properties of multiplication
+ **************************************)
+
+Theorem pmult_assoc: forall p q r, (pmult p (pmult q r)) = (pmult (pmult p q) r).
+intros p q r; case p; case q; case r; intros r1 r2 q1 q2 p1 p2; unfold pmult.
+eq_tac; ring.
+Qed.
+
+Theorem pmult_0_l: forall p, (pmult (0, 0) p) = (0, 0).
+intros p; case p; intros x y; unfold pmult; eq_tac; ring.
+Qed.
+
+Theorem pmult_0_r: forall p, (pmult p (0, 0)) = (0, 0).
+intros p; case p; intros x y; unfold pmult; eq_tac; ring.
+Qed.
+
+Theorem pmult_1_l: forall p, (pmult (1, 0) p) = p.
+intros p; case p; intros x y; unfold pmult; eq_tac; ring.
+Qed.
+
+Theorem pmult_1_r: forall p, (pmult p (1, 0)) = p.
+intros p; case p; intros x y; unfold pmult; eq_tac; ring.
+Qed.
+
+Theorem pmult_comm: forall p q, (pmult p q) = (pmult q p).
+intros p q; case p; case q; intros q1 q2 p1 p2; unfold pmult.
+eq_tac; ring.
+Qed.
+
+Theorem pplus_pmult_dist_l: forall p q r, (pmult p (pplus q r)) = (pplus (pmult p q) (pmult p r)).
+intros p q r; case p; case q; case r; intros r1 r2 q1 q2 p1 p2; unfold pplus, pmult.
+eq_tac; ring.
+Qed.
+
+
+Theorem pplus_pmult_dist_r: forall p q r, (pmult (pplus q r) p) = (pplus (pmult q p) (pmult r p)).
+intros p q r; case p; case q; case r; intros r1 r2 q1 q2 p1 p2; unfold pplus, pmult.
+eq_tac; ring.
+Qed.
+
+(**************************************
+ In this section we create the group PGroup of inversible elements {(p, q) | 0 <= p < m /\ 0 <= q < m}
+ **************************************)
+Section Mod.
+
+Variable m : Z.
+
+Hypothesis m_pos: 1 < m.
+
+(**************************************
+ mkLine creates {(a, p) | 0 <= p < n}
+ **************************************)
+
+Fixpoint mkLine (a: Z) (n: nat) {struct n} : list (Z * Z) :=
+ (a, Z_of_nat n) :: match n with O => nil | (S n1) => mkLine a n1 end.
+
+(**************************************
+ Some properties of mkLine
+ **************************************)
+
+Theorem mkLine_length: forall a n, length (mkLine a n) = (n + 1)%nat.
+intros a n; elim n; simpl; auto.
+Qed.
+
+Theorem mkLine_in: forall a n p, 0 <= p <= Z_of_nat n -> (In (a, p) (mkLine a n)).
+intros a n; elim n.
+simpl; auto with zarith.
+intros p (H1, H2); replace p with 0; auto with zarith.
+intros n1 Rec p (H1, H2).
+case (Zle_lt_or_eq p (Z_of_nat (S n1))); auto with zarith.
+rewrite inj_S in H2; auto with zarith.
+rewrite inj_S; auto with zarith.
+intros H3; right; apply Rec; auto with zarith.
+intros H3; subst; simpl; auto.
+Qed.
+
+Theorem in_mkLine: forall a n p, In p (mkLine a n) -> exists q, 0 <= q <= Z_of_nat n /\ p = (a, q).
+intros a n p; elim n; clear n.
+simpl; intros [H1 | H1]; exists 0; auto with zarith; case H1.
+simpl; intros n Rec [H1 | H1]; auto.
+exists (Z_of_nat (S n)); auto with zarith.
+case Rec; auto; intros q ((H2, H3), H4); exists q; repeat split; auto with zarith.
+change (q <= Z_of_nat (S n)).
+rewrite inj_S; auto with zarith.
+Qed.
+
+Theorem mkLine_ulist: forall a n, ulist (mkLine a n).
+intros a n; elim n; simpl; auto.
+intros n1 H; apply ulist_cons; auto.
+change (~ In (a, Z_of_nat (S n1)) (mkLine a n1)).
+rewrite inj_S; intros H1.
+case in_mkLine with (1 := H1); auto with zarith.
+intros x ((H2, H3), H4); injection H4.
+intros H5; subst; auto with zarith.
+Qed.
+
+(**************************************
+ mkRect creates the list {(p, q) | 0 <= p < n /\ 0 <= q < m}
+ **************************************)
+
+Fixpoint mkRect (n m: nat) {struct n} : list (Z * Z) :=
+ (mkLine (Z_of_nat n) m) ++ match n with O => nil | (S n1) => mkRect n1 m end.
+
+(**************************************
+ Some properties of mkRect
+ **************************************)
+
+Theorem mkRect_length: forall n m, length (mkRect n m) = ((n + 1) * (m + 1))%nat.
+intros n; elim n; simpl; auto.
+intros n1; rewrite <- app_nil_end; rewrite mkLine_length; rewrite plus_0_r; auto.
+intros n1 Rec m1; rewrite length_app; rewrite Rec; rewrite mkLine_length; auto.
+Qed.
+
+Theorem mkRect_in: forall n m p q, 0 <= p <= Z_of_nat n -> 0 <= q <= Z_of_nat m -> (In (p, q) (mkRect n m)).
+intros n m1; elim n; simpl.
+intros p q (H1, H2) (H3, H4); replace p with 0; auto with zarith.
+rewrite <- app_nil_end; apply mkLine_in; auto.
+intros n1 Rec p q (H1, H2) (H3, H4).
+case (Zle_lt_or_eq p (Z_of_nat (S n1))); auto with zarith; intros H5.
+rewrite inj_S in H5; apply in_or_app; auto with zarith.
+apply in_or_app; left; subst; apply mkLine_in; auto with zarith.
+Qed.
+
+Theorem in_mkRect: forall n m p, In p (mkRect n m) -> exists p1, exists p2, 0 <= p1 <= Z_of_nat n /\ 0 <= p2 <= Z_of_nat m /\ p = (p1, p2).
+intros n m1 p; elim n; clear n; simpl.
+rewrite <- app_nil_end; intros H1.
+case in_mkLine with (1 := H1).
+intros p2 (H2, H3); exists 0; exists p2; auto with zarith.
+intros n Rec H1.
+case in_app_or with (1 := H1); intros H2.
+case in_mkLine with (1 := H2).
+intros p2 (H3, H4); exists (Z_of_nat (S n)); exists p2; subst; simpl; auto with zarith.
+case Rec with (1 := H2); auto.
+intros p1 (p2, (H3, (H4, H5))); exists p1; exists p2; repeat split; auto with zarith.
+change (p1 <= Z_of_nat (S n)).
+rewrite inj_S; auto with zarith.
+Qed.
+
+Theorem mkRect_ulist: forall n m, ulist (mkRect n m).
+intros n; elim n; simpl; auto.
+intros n1; rewrite <- app_nil_end; apply mkLine_ulist; auto.
+intros n1 Rec m1; apply ulist_app; auto.
+apply mkLine_ulist.
+intros a H1 H2.
+case in_mkLine with (1 := H1); intros p1 ((H3, H4), H5).
+case in_mkRect with (1 := H2); intros p2 (p3, ((H6, H7), ((H8, H9), H10))).
+subst; injection H10; clear H10; intros; subst.
+contradict H7.
+change (~ Z_of_nat (S n1) <= Z_of_nat n1).
+rewrite inj_S; auto with zarith.
+Qed.
+
+(**************************************
+ mL is the list {(p, q) | 0 <= p < m-1 /\ 0 <= q < m - 1}
+ **************************************)
+Definition mL := mkRect (Zabs_nat (m - 1)) (Zabs_nat (m -1)).
+
+(**************************************
+ Some properties of mL
+ **************************************)
+
+Theorem mL_length : length mL = Zabs_nat (m * m).
+unfold mL; rewrite mkRect_length; simpl; apply inj_eq_rev.
+repeat (rewrite inj_mult || rewrite inj_plus || rewrite inj_Zabs_nat || rewrite Zabs_eq); simpl; auto with zarith.
+eq_tac; auto with zarith.
+Qed.
+
+Theorem mL_in: forall p q, 0 <= p < m -> 0 <= q < m -> (In (p, q) mL).
+intros p q (H1, H2) (H3, H4); unfold mL; apply mkRect_in; rewrite inj_Zabs_nat;
+ rewrite Zabs_eq; auto with zarith.
+Qed.
+
+Theorem in_mL: forall p, In p mL-> exists p1, exists p2, 0 <= p1 < m /\ 0 <= p2 < m /\ p = (p1, p2).
+unfold mL; intros p H1; case in_mkRect with (1 := H1).
+repeat (rewrite inj_Zabs_nat || rewrite Zabs_eq); auto with zarith.
+intros p1 (p2, ((H2, H3), ((H4, H5), H6))); exists p1; exists p2; repeat split; auto with zarith.
+Qed.
+
+Theorem mL_ulist: ulist mL.
+unfold mL; apply mkRect_ulist; auto.
+Qed.
+
+(**************************************
+ We define zpmult the multiplication of pairs module m
+ **************************************)
+
+Definition zpmult (p q: Z * Z) := let (x ,y) := pmult p q in (Zmod x m, Zmod y m).
+
+(**************************************
+ Some properties of zpmult
+ **************************************)
+
+Theorem zpmult_internal: forall p q, (In (zpmult p q) mL).
+intros p q; unfold zpmult; case (pmult p q); intros z y; apply mL_in; auto with zarith.
+apply Z_mod_lt; auto with zarith.
+apply Z_mod_lt; auto with zarith.
+Qed.
+
+Theorem zpmult_assoc: forall p q r, (zpmult p (zpmult q r)) = (zpmult (zpmult p q) r).
+assert (U: 0 < m); auto with zarith.
+intros p q r; unfold zpmult.
+generalize (pmult_assoc p q r).
+case (pmult p q); intros x1 x2.
+case (pmult q r); intros y1 y2.
+case p; case r; unfold pmult.
+intros z1 z2 t1 t2 H.
+match goal with
+ H: (?X, ?Y) = (?Z, ?T) |- _ =>
+ assert (H1: X = Z); assert (H2: Y = T); try (injection H; simpl; auto; fail); clear H
+end.
+eq_tac.
+generalize (f_equal (fun x => x mod m) H1).
+repeat rewrite <- Zmult_assoc.
+repeat (rewrite (fun x => Zplus_mod (t1 * x))); auto.
+repeat (rewrite (fun x => Zplus_mod (x1 * x))); auto.
+repeat (rewrite (fun x => Zplus_mod (x1 mod m * x))); auto.
+repeat (rewrite (Zmult_mod t1)); auto.
+repeat (rewrite (Zmult_mod x1)); auto.
+repeat (rewrite (Zmult_mod base)); auto.
+repeat (rewrite (Zmult_mod t2)); auto.
+repeat (rewrite (Zmult_mod x2)); auto.
+repeat (rewrite (Zmult_mod (t2 mod m))); auto.
+repeat (rewrite (Zmult_mod (x1 mod m))); auto.
+repeat (rewrite (Zmult_mod (x2 mod m))); auto.
+repeat (rewrite Zmod_mod); auto.
+generalize (f_equal (fun x => x mod m) H2).
+repeat (rewrite (fun x => Zplus_mod (t1 * x))); auto.
+repeat (rewrite (fun x => Zplus_mod (x1 * x))); auto.
+repeat (rewrite (fun x => Zplus_mod (x1 mod m * x))); auto.
+repeat (rewrite (Zmult_mod t1)); auto.
+repeat (rewrite (Zmult_mod x1)); auto.
+repeat (rewrite (Zmult_mod t2)); auto.
+repeat (rewrite (Zmult_mod x2)); auto.
+repeat (rewrite (Zmult_mod (t2 mod m))); auto.
+repeat (rewrite (Zmult_mod (x1 mod m))); auto.
+repeat (rewrite (Zmult_mod (x2 mod m))); auto.
+repeat (rewrite Zmod_mod); auto.
+Qed.
+
+Theorem zpmult_0_l: forall p, (zpmult (0, 0) p) = (0, 0).
+intros p; case p; intros x y; unfold zpmult, pmult; simpl.
+rewrite Zmod_small; auto with zarith.
+Qed.
+
+Theorem zpmult_1_l: forall p, In p mL -> zpmult (1, 0) p = p.
+intros p H; case in_mL with (1 := H); clear H; intros p1 (p2, ((H1, H2), (H3, H4))); subst.
+unfold zpmult; rewrite pmult_1_l.
+repeat rewrite Zmod_small; auto with zarith.
+Qed.
+
+Theorem zpmult_1_r: forall p, In p mL -> zpmult p (1, 0) = p.
+intros p H; case in_mL with (1 := H); clear H; intros p1 (p2, ((H1, H2), (H3, H4))); subst.
+unfold zpmult; rewrite pmult_1_r.
+repeat rewrite Zmod_small; auto with zarith.
+Qed.
+
+Theorem zpmult_comm: forall p q, zpmult p q = zpmult q p.
+intros p q; unfold zpmult; rewrite pmult_comm; auto.
+Qed.
+
+(**************************************
+ We are now ready to build our group
+ **************************************)
+
+Definition PGroup : (FGroup zpmult).
+apply IGroup with (support := mL) (e:= (1, 0)).
+exact P_dec.
+apply mL_ulist.
+apply mL_in; auto with zarith.
+intros; apply zpmult_internal.
+intros; apply zpmult_assoc.
+exact zpmult_1_l.
+exact zpmult_1_r.
+Defined.
+
+End Mod.
diff --git a/coqprime/PrimalityTest/Pepin.v b/coqprime/PrimalityTest/Pepin.v
new file mode 100644
index 000000000..c400e0a43
--- /dev/null
+++ b/coqprime/PrimalityTest/Pepin.v
@@ -0,0 +1,123 @@
+
+(*************************************************************)
+(* This file is distributed under the terms of the *)
+(* GNU Lesser General Public License Version 2.1 *)
+(*************************************************************)
+(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
+(*************************************************************)
+
+(**********************************************************************
+ Pepin.v
+
+ Pepin's Test for Fermat Number
+
+ Definition: PepinTest
+ **********************************************************************)
+Require Import ZArith.
+Require Import ZCAux.
+Require Import Pocklington.
+
+Open Scope Z_scope.
+
+Definition FermatNumber n := 2^(2^(Z_of_nat n)) + 1.
+
+Theorem Fermat_pos: forall n, 1 < FermatNumber n.
+unfold FermatNumber; intros n; apply Zle_lt_trans with (2 ^ 2 ^(Z_of_nat n)); auto with zarith.
+rewrite <- (Zpower_0_r 2); auto with zarith.
+apply Zpower_le_monotone; try split; auto with zarith.
+Qed.
+
+Theorem PepinTest: forall n, let Fn := FermatNumber n in (3 ^ ((Fn - 1) / 2) + 1) mod Fn = 0 -> prime Fn.
+intros n Fn H.
+assert (Hn: 1 < Fn).
+unfold Fn; apply Fermat_pos.
+apply PocklingtonCorollary1 with (F1 := 2^(2^(Z_of_nat n))) (R1 := 1); auto with zarith.
+2: unfold Fn, FermatNumber; auto with zarith.
+apply Zlt_le_trans with (2 ^ 1); auto with zarith.
+rewrite Zpower_1_r; auto with zarith.
+apply Zpower_le_monotone; try split; auto with zarith.
+rewrite <- (Zpower_0_r 2); apply Zpower_le_monotone; try split; auto with zarith.
+unfold Fn, FermatNumber.
+assert (H1: 2 <= 2 ^ 2 ^ Z_of_nat n).
+pattern 2 at 1; rewrite <- (Zpower_1_r 2); auto with zarith.
+apply Zpower_le_monotone; split; auto with zarith.
+rewrite <- (Zpower_0_r 2); apply Zpower_le_monotone; try split; auto with zarith.
+apply Zlt_le_trans with (2 * 2 ^2 ^Z_of_nat n).
+assert (tmp: forall p, 2 * p = p + p); auto with zarith.
+apply Zmult_le_compat_r; auto with zarith.
+assert (Hd: (2 | Fn - 1)).
+exists (2 ^ (2^(Z_of_nat n) - 1)).
+pattern 2 at 3; rewrite <- (Zpower_1_r 2).
+rewrite <- Zpower_exp; auto with zarith.
+assert (tmp: forall p, p = (p - 1) +1); auto with zarith; rewrite <- tmp.
+unfold Fn, FermatNumber; ring.
+assert (0 < 2 ^ Z_of_nat n); auto with zarith.
+intros p Hp Hp1; exists 3; split; auto with zarith; split; auto.
+rewrite (Zdivide_Zdiv_eq 2 (Fn -1)); auto with zarith.
+rewrite Zmult_comm; rewrite Zpower_mult; auto with zarith.
+rewrite Zpower_mod; auto with zarith.
+assert (tmp: forall p, p = (p + 1) -1); auto with zarith; rewrite (fun x => (tmp (3 ^ x))).
+rewrite Zminus_mod; auto with zarith.
+rewrite H.
+rewrite (Zmod_small 1); auto with zarith.
+rewrite <- Zpower_mod; auto with zarith.
+rewrite Zmod_small; auto with zarith.
+simpl; unfold Zpower_pos; simpl; auto with zarith.
+apply Z_div_pos; auto with zarith.
+apply Zis_gcd_gcd; auto with zarith.
+apply Zis_gcd_intro; auto with zarith.
+intros x HD1 HD2.
+assert (Hd1: p = 2).
+apply prime_div_Zpower_prime with (4 := Hp1); auto with zarith.
+apply prime_2.
+assert (Hd2: (x | 2)).
+replace 2 with ((3 ^ ((Fn - 1) / 2) + 1) - (3 ^ ((Fn - 1) / 2) - 1)); auto with zarith.
+apply Zdivide_minus_l; auto.
+apply Zdivide_trans with (1 := HD2).
+apply Zmod_divide; auto with zarith.
+rewrite <- Hd1; auto.
+replace 1 with (Fn - (Fn - 1)); auto with zarith.
+apply Zdivide_minus_l; auto.
+apply Zdivide_trans with (1 := Hd2); auto.
+Qed.
+
+(* An optimized version with Zpow_mod *)
+
+Definition pepin_test n :=
+ let Fn := FermatNumber n in if Z_eq_dec (Zpow_mod 3 ((Fn - 1) / 2) Fn) (Fn - 1) then true else false.
+
+Theorem PepinTestOp: forall n, pepin_test n = true -> prime (FermatNumber n).
+intros n; unfold pepin_test.
+match goal with |- context[if ?X then _ else _] => case X end; try (intros; discriminate).
+intros H1 _; apply PepinTest.
+generalize (Fermat_pos n); intros H2.
+rewrite Zplus_mod; auto with zarith.
+rewrite <- Zpow_mod_Zpower_correct; auto with zarith.
+rewrite H1.
+rewrite (Zmod_small 1); auto with zarith.
+replace (FermatNumber n - 1 + 1) with (FermatNumber n); auto with zarith.
+apply Zdivide_mod; auto with zarith.
+apply Z_div_pos; auto with zarith.
+Qed.
+
+Theorem prime5: prime 5.
+exact (PepinTestOp 1 (refl_equal _)).
+Qed.
+
+Theorem prime17: prime 17.
+exact (PepinTestOp 2 (refl_equal _)).
+Qed.
+
+Theorem prime257: prime 257.
+exact (PepinTestOp 3 (refl_equal _)).
+Qed.
+
+Theorem prime65537: prime 65537.
+exact (PepinTestOp 4 (refl_equal _)).
+Qed.
+
+(* Too tough !!
+Theorem prime4294967297: prime 4294967297.
+refine (PepinTestOp 5 (refl_equal _)).
+Qed.
+*)
diff --git a/coqprime/PrimalityTest/Pocklington.v b/coqprime/PrimalityTest/Pocklington.v
new file mode 100644
index 000000000..9871cd3e6
--- /dev/null
+++ b/coqprime/PrimalityTest/Pocklington.v
@@ -0,0 +1,261 @@
+
+(*************************************************************)
+(* This file is distributed under the terms of the *)
+(* GNU Lesser General Public License Version 2.1 *)
+(*************************************************************)
+(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
+(*************************************************************)
+
+Require Import ZArith.
+Require Export Znumtheory.
+Require Import Tactic.
+Require Import ZCAux.
+Require Import Zp.
+Require Import FGroup.
+Require Import EGroup.
+Require Import Euler.
+
+Open Scope Z_scope.
+
+Theorem Pocklington:
+forall N F1 R1, 1 < F1 -> 0 < R1 -> N - 1 = F1 * R1 ->
+ (forall p, prime p -> (p | F1) -> exists a, 1 < a /\ a^(N - 1) mod N = 1 /\ Zgcd (a ^ ((N - 1)/ p) - 1) N = 1) ->
+ forall n, prime n -> (n | N) -> n mod F1 = 1.
+intros N F1 R1 HF1 HR1 Neq Rec n Hn H.
+assert (HN: 1 < N).
+assert (0 < N - 1); auto with zarith.
+rewrite Neq; auto with zarith.
+apply Zlt_le_trans with (1* R1); auto with zarith.
+assert (Hn1: 1 < n); auto with zarith.
+apply Zlt_le_trans with 2; try apply prime_ge_2; auto with zarith.
+assert (H1: (F1 | n - 1)).
+2: rewrite <- (Zmod_small 1 F1); auto with zarith.
+2: case H1; intros k H1'.
+2: replace n with (1 + (n - 1)); auto with zarith.
+2: rewrite H1'; apply Z_mod_plus; auto with zarith.
+apply Zdivide_Zpower; auto with zarith.
+intros p i Hp Hi HiF1.
+case (Rec p); auto.
+apply Zdivide_trans with (2 := HiF1).
+apply Zpower_divide; auto with zarith.
+intros a (Ha1, (Ha2, Ha3)).
+assert (HNn: a ^ (N - 1) mod n = 1).
+apply Zdivide_mod_minus; auto with zarith.
+apply Zdivide_trans with (1 := H).
+apply Zmod_divide_minus; auto with zarith.
+assert (~(n | a)).
+intros H1; absurd (0 = 1); auto with zarith.
+rewrite <- HNn; auto.
+apply sym_equal; apply Zdivide_mod; auto with zarith.
+apply Zdivide_trans with (1 := H1); apply Zpower_divide; auto with zarith.
+assert (Hr: rel_prime a n).
+apply rel_prime_sym; apply prime_rel_prime; auto.
+assert (Hz: 0 < Zorder a n).
+apply Zorder_power_pos; auto.
+apply Zdivide_trans with (Zorder a n).
+apply prime_divide_Zpower_Zdiv with (N - 1); auto with zarith.
+apply Zorder_div_power; auto with zarith.
+intros H1; absurd (1 < n); auto; apply Zle_not_lt; apply Zdivide_le; auto with zarith.
+rewrite <- Ha3; apply Zdivide_Zgcd; auto with zarith.
+apply Zmod_divide_minus; auto with zarith.
+case H1; intros t Ht; rewrite Ht.
+assert (Ht1: 0 <= t).
+apply Zmult_le_reg_r with (Zorder a n); auto with zarith.
+rewrite Zmult_0_l; rewrite <- Ht.
+apply Zge_le; apply Z_div_ge0; auto with zarith.
+apply Zlt_gt; apply Zlt_le_trans with 2; try apply prime_ge_2; auto with zarith.
+rewrite Zmult_comm; rewrite Zpower_mult; auto with zarith.
+rewrite Zpower_mod; auto with zarith.
+rewrite Zorder_power_is_1; auto with zarith.
+rewrite Zpower_1_l; auto with zarith.
+apply Zmod_small; auto with zarith.
+apply Zdivide_trans with (1:= HiF1); rewrite Neq; apply Zdivide_factor_r.
+apply Zorder_div; auto.
+Qed.
+
+Theorem PocklingtonCorollary1:
+forall N F1 R1, 1 < F1 -> 0 < R1 -> N - 1 = F1 * R1 -> N < F1 * F1 ->
+ (forall p, prime p -> (p | F1) -> exists a, 1 < a /\ a^(N - 1) mod N = 1 /\ Zgcd (a ^ ((N - 1)/ p) - 1) N = 1) ->
+ prime N.
+intros N F1 R1 H H1 H2 H3 H4; case (prime_dec N); intros H5; auto.
+assert (HN: 1 < N).
+assert (0 < N - 1); auto with zarith.
+rewrite H2; auto with zarith.
+apply Zlt_le_trans with (1* R1); auto with zarith.
+case Zdivide_div_prime_le_square with (2:= H5); auto with zarith.
+intros n (Hn, (Hn1, Hn2)).
+assert (Hn3: 0 <= n).
+apply Zle_trans with 2; try apply prime_ge_2; auto with zarith.
+absurd (n = 1).
+intros H6; contradict Hn; subst; apply not_prime_1.
+rewrite <- (Zmod_small n F1); try split; auto.
+apply Pocklington with (R1 := R1) (4 := H4); auto.
+apply Zlt_square_mult_inv; auto with zarith.
+Qed.
+
+Theorem PocklingtonCorollary2:
+forall N F1 R1, 1 < F1 -> 0 < R1 -> N - 1 = F1 * R1 ->
+ (forall p, prime p -> (p | F1) -> exists a, 1 < a /\ a^(N - 1) mod N = 1 /\ Zgcd (a ^ ((N - 1)/ p) - 1) N = 1) ->
+ forall n, 0 <= n -> (n | N) -> n mod F1 = 1.
+intros N F1 R1 H1 H2 H3 H4 n H5; pattern n; apply prime_induction; auto.
+assert (HN: 1 < N).
+assert (0 < N - 1); auto with zarith.
+rewrite H3; auto with zarith.
+apply Zlt_le_trans with (1* R1); auto with zarith.
+intros (u, Hu); contradict HN; subst; rewrite Zmult_0_r; auto with zarith.
+intro H6; rewrite Zmod_small; auto with zarith.
+intros p q Hp Hp1 Hp2; rewrite Zmult_mod; auto with zarith.
+rewrite Pocklington with (n := p) (R1 := R1) (4 := H4); auto.
+rewrite Hp1.
+rewrite Zmult_1_r; rewrite Zmod_small; auto with zarith.
+apply Zdivide_trans with (2 := Hp2); apply Zdivide_factor_l.
+apply Zdivide_trans with (2 := Hp2); apply Zdivide_factor_r; auto.
+Qed.
+
+Definition isSquare x := exists y, x = y * y.
+
+Theorem PocklingtonExtra:
+forall N F1 R1, 1 < F1 -> 0 < R1 -> N - 1 = F1 * R1 -> Zeven F1 -> Zodd R1 ->
+ (forall p, prime p -> (p | F1) -> exists a, 1 < a /\ a^(N - 1) mod N = 1 /\ Zgcd (a ^ ((N - 1)/ p) - 1) N = 1) ->
+ forall m, 1 <= m -> (forall l, 1 <= l < m -> ~((l * F1 + 1) | N)) ->
+ let s := (R1 / (2 * F1)) in
+ let r := (R1 mod (2 * F1)) in
+ N < (m * F1 + 1) * (2 * F1 * F1 + (r - m) * F1 + 1) ->
+ (s = 0 \/ ~ isSquare (r * r - 8 * s)) -> prime N.
+intros N F1 R1 H1 H2 H3 OF1 ER1 H4 m H5 H6 s r H7 H8.
+case (prime_dec N); auto; intros H9.
+assert (HN: 1 < N).
+assert (0 < N - 1); auto with zarith.
+rewrite H3; auto with zarith.
+apply Zlt_le_trans with (1* R1); auto with zarith.
+case Zdivide_div_prime_le_square with N; auto.
+intros X (Hx1, (Hx2, Hx3)).
+assert (Hx0: 1 < X).
+apply Zlt_le_trans with 2; try apply prime_ge_2; auto with zarith.
+pose (c := (X / F1)).
+assert(Hc1: 0 <= c); auto with zarith.
+apply Zge_le; unfold c; apply Z_div_ge0; auto with zarith.
+assert (Hc2: X = c * F1 + 1).
+rewrite (Z_div_mod_eq X F1); auto with zarith.
+eq_tac; auto.
+rewrite (Zmult_comm F1); auto.
+apply PocklingtonCorollary2 with (R1 := R1) (4 := H4); auto with zarith.
+case Zle_lt_or_eq with (1 := Hc1); clear Hc1; intros Hc1.
+2: contradict Hx0; rewrite Hc2; try rewrite <- Hc1; auto with zarith.
+case (Zle_or_lt m c); intros Hc3.
+2: case Zle_lt_or_eq with (1 := H5); clear H5; intros H5; auto with zarith.
+2: case (H6 c); auto with zarith; rewrite <- Hc2; auto.
+2: contradict Hc3; rewrite <- H5; auto with zarith.
+pose (d := ((N / X) / F1)).
+assert(Hd0: 0 <= N / X); try apply Z_div_pos; auto with zarith.
+(*
+apply Zge_le; unfold d; repeat apply Z_div_ge0; auto with zarith.
+*)
+assert(Hd1: 0 <= d); auto with zarith.
+apply Zge_le; unfold d; repeat apply Z_div_ge0; auto with zarith.
+assert (Hd2: N / X = d * F1 + 1).
+rewrite (Z_div_mod_eq (N / X) F1); auto with zarith.
+eq_tac; auto.
+rewrite (Zmult_comm F1); auto.
+apply PocklingtonCorollary2 with (R1 := R1) (4 := H4); auto with zarith.
+exists X; auto with zarith.
+apply Zdivide_Zdiv_eq; auto with zarith.
+case Zle_lt_or_eq with (1 := Hd0); clear Hd0; intros Hd0.
+2: contradict HN; rewrite (Zdivide_Zdiv_eq X N); auto with zarith.
+2: rewrite <- Hd0; auto with zarith.
+case (Zle_lt_or_eq 1 (N / X)); auto with zarith; clear Hd0; intros Hd0.
+2: contradict H9; rewrite (Zdivide_Zdiv_eq X N); auto with zarith.
+2: rewrite <- Hd0; rewrite Zmult_1_r; auto with zarith.
+case Zle_lt_or_eq with (1 := Hd1); clear Hd1; intros Hd1.
+2: contradict Hd0; rewrite Hd2; try rewrite <- Hd1; auto with zarith.
+case (Zle_or_lt m d); intros Hd3.
+2: case Zle_lt_or_eq with (1 := H5); clear H5; intros H5; auto with zarith.
+2: case (H6 d); auto with zarith; rewrite <- Hd2; auto.
+2: exists X; auto with zarith.
+2: apply Zdivide_Zdiv_eq; auto with zarith.
+2: contradict Hd3; rewrite <- H5; auto with zarith.
+assert (L5: N = (c * F1 + 1) * (d * F1 + 1)).
+rewrite <- Hc2; rewrite <- Hd2; apply Zdivide_Zdiv_eq; auto with zarith.
+assert (L6: R1 = c * d * F1 + c + d).
+apply trans_equal with ((N - 1) / F1).
+rewrite H3; rewrite Zmult_comm; apply sym_equal; apply Z_div_mult; auto with zarith.
+rewrite L5.
+match goal with |- (?X / ?Y = ?Z) => replace X with (Z * Y) end; try ring; apply Z_div_mult; auto with zarith.
+assert (L6_1: Zodd (c + d)).
+case (Zeven_odd_dec (c + d)); auto; intros O1.
+contradict ER1; apply Zeven_not_Zodd; rewrite L6; rewrite <- Zplus_assoc; apply Zeven_plus_Zeven; auto.
+apply Zeven_mult_Zeven_r; auto.
+assert (L6_2: Zeven (c * d)).
+case (Zeven_odd_dec c); intros HH1.
+apply Zeven_mult_Zeven_l; auto.
+case (Zeven_odd_dec d); intros HH2.
+apply Zeven_mult_Zeven_r; auto.
+contradict L6_1; apply Zeven_not_Zodd; apply Zodd_plus_Zodd; auto.
+assert ((c + d) mod (2 * F1) = r).
+rewrite <- Z_mod_plus with (b := Zdiv2 (c * d)); auto with zarith.
+match goal with |- ?X mod _ = _ => replace X with R1 end; auto.
+rewrite L6; pattern (c * d) at 1.
+rewrite Zeven_div2 with (1 := L6_2); ring.
+assert (L9: c + d - r < 2 * F1).
+apply Zplus_lt_reg_r with (r - m).
+apply Zmult_lt_reg_r with (F1); auto with zarith.
+apply Zplus_lt_reg_r with 1.
+match goal with |- ?X < ?Y =>
+ replace Y with (2 * F1 * F1 + (r - m) * F1 + 1); try ring;
+ replace X with ((((c + d) - m) * F1) + 1); try ring
+end.
+apply Zmult_lt_reg_r with (m * F1 + 1); auto with zarith.
+apply Zlt_trans with (m * F1 + 0); auto with zarith.
+rewrite Zplus_0_r; apply Zmult_lt_O_compat; auto with zarith.
+repeat rewrite (fun x => Zmult_comm x (m * F1 + 1)).
+apply Zle_lt_trans with (2 := H7).
+rewrite L5.
+match goal with |- ?X <= ?Y =>
+ replace X with ((m * (c + d) - m * m ) * F1 * F1 + (c + d) * F1 + 1); try ring;
+ replace Y with ((c * d) * F1 * F1 + (c + d) * F1 + 1); try ring
+end.
+repeat apply Zplus_le_compat_r.
+repeat apply Zmult_le_compat_r; auto with zarith.
+assert (tmp: forall p q, 0 <= p - q -> q <= p); auto with zarith; try apply tmp.
+match goal with |- _ <= ?X =>
+ replace X with ((c - m) * (d - m)); try ring; auto with zarith
+end.
+assert (L10: c + d = r).
+apply Zmod_closeby_eq with (2 * F1); auto with zarith.
+unfold r; apply Z_mod_lt; auto with zarith.
+assert (L11: 2 * s = c * d).
+apply Zmult_reg_r with F1; auto with zarith.
+apply trans_equal with (R1 - (c + d)).
+rewrite L10; rewrite (Z_div_mod_eq R1 (2 * F1)); auto with zarith.
+unfold s, r; ring.
+rewrite L6; ring.
+case H8; intro H10.
+absurd (0 < c * d); auto with zarith.
+apply Zmult_lt_O_compat; auto with zarith.
+case H10; exists (c - d); auto with zarith.
+rewrite <- L10.
+replace (8 * s) with (4 * (2 * s)); auto with zarith; try rewrite L11; ring.
+Qed.
+
+Theorem PocklingtonExtraCorollary:
+forall N F1 R1, 1 < F1 -> 0 < R1 -> N - 1 = F1 * R1 -> Zeven F1 -> Zodd R1 ->
+ (forall p, prime p -> (p | F1) -> exists a, 1 < a /\ a^(N - 1) mod N = 1 /\ Zgcd (a ^ ((N - 1)/ p) - 1) N = 1) ->
+ let s := (R1 / (2 * F1)) in
+ let r := (R1 mod (2 * F1)) in
+ N < 2 * F1 * F1 * F1 -> (s = 0 \/ ~ isSquare (r * r - 8 * s)) -> prime N.
+intros N F1 R1 H1 H2 H3 OF1 ER1 H4 s r H5 H6.
+apply PocklingtonExtra with (6 := H4) (R1 := R1) (m := 1); auto with zarith.
+apply Zlt_le_trans with (1 := H5).
+match goal with |- ?X <= ?K * ((?Y + ?Z) + ?T) =>
+ rewrite <- (Zplus_0_l X);
+ replace (K * ((Y + Z) + T)) with ((F1 * (Z + T) + Y + Z + T) + X);[idtac | ring]
+end.
+apply Zplus_le_compat_r.
+case (Zle_lt_or_eq 0 r); unfold r; auto with zarith.
+case (Z_mod_lt R1 (2 * F1)); auto with zarith.
+intros HH; repeat ((rewrite <- (Zplus_0_r 0); apply Zplus_le_compat)); auto with zarith.
+intros HH; contradict ER1; apply Zeven_not_Zodd.
+rewrite (Z_div_mod_eq R1 (2 * F1)); auto with zarith.
+rewrite <- HH; rewrite Zplus_0_r.
+rewrite <- Zmult_assoc; apply Zeven_2p.
+Qed.
diff --git a/coqprime/PrimalityTest/PocklingtonCertificat.v b/coqprime/PrimalityTest/PocklingtonCertificat.v
new file mode 100644
index 000000000..ed75ca281
--- /dev/null
+++ b/coqprime/PrimalityTest/PocklingtonCertificat.v
@@ -0,0 +1,759 @@
+
+(*************************************************************)
+(* This file is distributed under the terms of the *)
+(* GNU Lesser General Public License Version 2.1 *)
+(*************************************************************)
+(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
+(*************************************************************)
+
+Require Import List.
+Require Import ZArith.
+Require Import Zorder.
+Require Import ZCAux.
+Require Import LucasLehmer.
+Require Import Pocklington.
+Require Import ZCmisc.
+Require Import Pmod.
+
+Definition dec_prime := list (positive * positive).
+
+Inductive singleCertif : Set :=
+ | Proof_certif : forall N:positive, prime N -> singleCertif
+ | Lucas_certif : forall (n:positive) (p: Z), singleCertif
+ | Pock_certif : forall N a : positive, dec_prime -> positive -> singleCertif
+ | SPock_certif : forall N a : positive, dec_prime -> singleCertif
+ | Ell_certif: forall (N S: positive) (l: list (positive * positive))
+ (A B x y: Z), singleCertif.
+
+Definition Certif := list singleCertif.
+
+Definition nprim sc :=
+ match sc with
+ | Proof_certif n _ => n
+ | Lucas_certif n _ => n
+ | Pock_certif n _ _ _ => n
+ | SPock_certif n _ _ => n
+ | Ell_certif n _ _ _ _ _ _ => n
+
+ end.
+
+Open Scope positive_scope.
+Open Scope P_scope.
+
+Fixpoint pow (a p:positive) {struct p} : positive :=
+ match p with
+ | xH => a
+ | xO p' =>let z := pow a p' in square z
+ | xI p' => let z := pow a p' in square z * a
+ end.
+
+Definition mkProd' (l:dec_prime) :=
+ fold_right (fun (k:positive*positive) r => times (fst k) r) 1%positive l.
+
+Definition mkProd_pred (l:dec_prime) :=
+ fold_right (fun (k:positive*positive) r =>
+ if ((snd k) ?= 1)%P then r else times (pow (fst k) (Ppred (snd k))) r)
+ 1%positive l.
+
+Definition mkProd (l:dec_prime) :=
+ fold_right (fun (k:positive*positive) r => times (pow (fst k) (snd k)) r) 1%positive l.
+
+(* [pow_mod a m n] return [a^m mod n] *)
+Fixpoint pow_mod (a m n : positive) {struct m} : N :=
+ match m with
+ | xH => (a mod n)
+ | xO m' =>
+ let z := pow_mod a m' n in
+ match z with
+ | N0 => 0%N
+ | Npos z' => ((square z') mod n)
+ end
+ | xI m' =>
+ let z := pow_mod a m' n in
+ match z with
+ | N0 => 0%N
+ | Npos z' => ((square z') * a)%P mod n
+ end
+ end.
+
+Definition Npow_mod a m n :=
+ match a with
+ | N0 => 0%N
+ | Npos a => pow_mod a m n
+ end.
+
+(* [fold_pow_mod a [q1,_;...;qn,_]] b = a ^(q1*...*qn) mod b *)
+(* invariant a mod N = a *)
+Definition fold_pow_mod a l n :=
+ fold_left
+ (fun a' (qp:positive*positive) => Npow_mod a' (fst qp) n)
+ l a.
+
+Definition times_mod x y n :=
+ match x, y with
+ | N0, _ => N0
+ | _, N0 => N0
+ | Npos x, Npos y => ((x * y)%P mod n)
+ end.
+
+Definition Npred_mod p n :=
+ match p with
+ | N0 => Npos (Ppred n)
+ | Npos p =>
+ if (p ?= 1) then N0
+ else Npos (Ppred p)
+ end.
+
+Fixpoint all_pow_mod (prod a : N) (l:dec_prime) (n:positive) {struct l}: N*N :=
+ match l with
+ | nil => (prod,a)
+ | (q,_) :: l =>
+ let m := Npred_mod (fold_pow_mod a l n) n in
+ all_pow_mod (times_mod prod m n) (Npow_mod a q n) l n
+ end.
+
+Fixpoint pow_mod_pred (a:N) (l:dec_prime) (n:positive) {struct l} : N :=
+ match l with
+ | nil => a
+ | (q,p)::l =>
+ if (p ?= 1) then pow_mod_pred a l n
+ else
+ let a' := iter_pos (Ppred p) _ (fun x => Npow_mod x q n) a in
+ pow_mod_pred a' l n
+ end.
+
+Definition is_odd p :=
+ match p with
+ | xO _ => false
+ | _ => true
+ end.
+
+Definition is_even p :=
+ match p with
+ | xO _ => true
+ | _ => false
+ end.
+
+Definition check_s_r s r sqrt :=
+ match s with
+ | N0 => true
+ | Npos p =>
+ match (Zminus (square r) (xO (xO (xO p)))) with
+ | Zpos x =>
+ let sqrt2 := square sqrt in
+ let sqrt12 := square (Psucc sqrt) in
+ if sqrt2 ?< x then x ?< sqrt12
+ else false
+ | Zneg _ => true
+ | Z0 => false
+ end
+ end.
+
+Definition test_pock N a dec sqrt :=
+ if (2 ?< N) then
+ let Nm1 := Ppred N in
+ let F1 := mkProd dec in
+ match Nm1 / F1 with
+ | (Npos R1, N0) =>
+ if is_odd R1 then
+ if is_even F1 then
+ if (1 ?< a) then
+ let (s,r') := (R1 / (xO F1))in
+ match r' with
+ | Npos r =>
+ let A := pow_mod_pred (pow_mod a R1 N) dec N in
+ match all_pow_mod 1%N A dec N with
+ | (Npos p, Npos aNm1) =>
+ if (aNm1 ?= 1) then
+ if gcd p N ?= 1 then
+ if check_s_r s r sqrt then
+ (N ?< (times ((times ((xO F1)+r+1) F1) + r) F1) + 1)
+ else false
+ else false
+ else false
+ | _ => false
+ end
+ | _ => false
+ end
+ else false
+ else false
+ else false
+ | _=> false
+ end
+ else false.
+
+Fixpoint is_in (p : positive) (lc : Certif) {struct lc} : bool :=
+ match lc with
+ | nil => false
+ | c :: l => if p ?= (nprim c) then true else is_in p l
+ end.
+
+Fixpoint all_in (lc : Certif) (lp : dec_prime) {struct lp} : bool :=
+ match lp with
+ | nil => true
+ | (p,_) :: lp =>
+ if all_in lc lp
+ then is_in p lc
+ else false
+ end.
+
+Definition gt2 n :=
+ match n with
+ | Zpos p => (2 ?< p)%positive
+ | _ => false
+ end.
+
+Fixpoint test_Certif (lc : Certif) : bool :=
+ match lc with
+ | nil => true
+ | (Proof_certif _ _) :: lc => test_Certif lc
+ | (Lucas_certif n p) :: lc =>
+ if test_Certif lc then
+ if gt2 p then
+ match Mp p with
+ | Zpos n' =>
+ if (n ?= n') then
+ match SS p with
+ | Z0 => true
+ | _ => false
+ end
+ else false
+ | _ => false
+ end
+ else false
+ else false
+ | (Pock_certif n a dec sqrt) :: lc =>
+ if test_pock n a dec sqrt then
+ if all_in lc dec then test_Certif lc else false
+ else false
+(* Shoudl be done later to do it with Z *)
+ | (SPock_certif n a dec) :: lc => false
+ | (Ell_certif _ _ _ _ _ _ _):: lc => false
+ end.
+
+Lemma pos_eq_1_spec :
+ forall p,
+ if (p ?= 1)%P then p = xH
+ else (1 < p).
+Proof.
+ unfold Zlt;destruct p;simpl; auto; red;reflexivity.
+Qed.
+
+Open Scope Z_scope.
+Lemma mod_unique : forall b q1 r1 q2 r2,
+ 0 <= r1 < b ->
+ 0 <= r2 < b ->
+ b * q1 + r1 = b * q2 + r2 ->
+ q1 = q2 /\ r1 = r2.
+Proof with auto with zarith.
+ intros b q1 r1 q2 r2 H1 H2 H3.
+ assert (r2 = (b * q1 + r1) -b*q2). rewrite H3;ring.
+ assert (b*(q2 - q1) = r1 - r2 ). rewrite H;ring.
+ assert (-b < r1 - r2 < b). omega.
+ destruct (Ztrichotomy q1 q2) as [H5 | [H5 | H5]].
+ assert (q2 - q1 >= 1). omega.
+ assert (r1- r2 >= b).
+ rewrite <- H0.
+ pattern b at 2; replace b with (b*1).
+ apply Zmult_ge_compat_l; omega. ring.
+ elimtype False; omega.
+ split;trivial. rewrite H;rewrite H5;ring.
+ assert (r1- r2 <= -b).
+ rewrite <- H0.
+ replace (-b) with (b*(-1)); try (ring;fail).
+ apply Zmult_le_compat_l; omega.
+ elimtype False; omega.
+Qed.
+
+Lemma Zge_0_pos : forall p:positive, p>= 0.
+Proof.
+ intros;unfold Zge;simpl;intro;discriminate.
+Qed.
+
+Lemma Zge_0_pos_add : forall p:positive, p+p>= 0.
+Proof.
+ intros;simpl;apply Zge_0_pos.
+Qed.
+
+Hint Resolve Zpower_gt_0 Zlt_0_pos Zge_0_pos Zlt_le_weak Zge_0_pos_add: zmisc.
+
+Hint Rewrite Zpos_mult Zpower_mult Zpower_1_r Zmod_mod Zpower_exp
+ times_Zmult square_Zmult Psucc_Zplus: zmisc.
+
+Ltac mauto :=
+ trivial;autorewrite with zmisc;trivial;auto with zmisc zarith.
+
+Lemma mod_lt : forall a (b:positive), a mod b < b.
+Proof.
+ intros a b;destruct (Z_mod_lt a b);mauto.
+Qed.
+Hint Resolve mod_lt : zmisc.
+
+Lemma Zmult_mod_l : forall (n:positive) a b, (a mod n * b) mod n = (a * b) mod n.
+Proof with mauto.
+ intros;rewrite Zmult_mod ... rewrite (Zmult_mod a) ...
+Qed.
+
+Lemma Zmult_mod_r : forall (n:positive) a b, (a * (b mod n)) mod n = (a * b) mod n.
+Proof with mauto.
+ intros;rewrite Zmult_mod ... rewrite (Zmult_mod a) ...
+Qed.
+
+Lemma Zminus_mod_l : forall (n:positive) a b, (a mod n - b) mod n = (a - b) mod n.
+Proof with mauto.
+ intros;rewrite Zminus_mod ... rewrite (Zminus_mod a) ...
+Qed.
+
+Lemma Zminus_mod_r : forall (n:positive) a b, (a - (b mod n)) mod n = (a - b) mod n.
+Proof with mauto.
+ intros;rewrite Zminus_mod ... rewrite (Zminus_mod a) ...
+Qed.
+
+Hint Rewrite Zmult_mod_l Zmult_mod_r Zminus_mod_l Zminus_mod_r : zmisc.
+Hint Rewrite <- Zpower_mod : zmisc.
+
+Lemma Pmod_Zmod : forall a b, Z_of_N (a mod b)%P = a mod b.
+Proof.
+ intros a b; rewrite Pmod_div_eucl.
+ assert (b>0). mauto.
+ unfold Zmod; assert (H1 := Z_div_mod a b H).
+ destruct (Zdiv_eucl a b) as (q2, r2).
+ assert (H2 := div_eucl_spec a b).
+ assert (Z_of_N (fst (a / b)%P) = q2 /\ Z_of_N (snd (a/b)%P) = r2).
+ destruct H1;destruct H2.
+ apply mod_unique with b;mauto.
+ split;mauto.
+ unfold Zle;destruct (snd (a / b)%P);intro;discriminate.
+ rewrite <- H0;symmetry;rewrite Zmult_comm;trivial.
+ destruct H0;auto.
+Qed.
+Hint Rewrite Pmod_Zmod : zmisc.
+
+Lemma Zpower_0 : forall p : positive, 0^p = 0.
+Proof.
+ intros;simpl;destruct p;unfold Zpower_pos;simpl;trivial.
+ generalize (iter_pos p Z (Z.mul 0) 1).
+ induction p;simpl;trivial.
+Qed.
+
+Opaque Zpower.
+Opaque Zmult.
+
+Lemma pow_Zpower : forall a p, Zpos (pow a p) = a ^ p.
+Proof with mauto.
+ induction p;simpl... rewrite IHp... rewrite IHp...
+Qed.
+Hint Rewrite pow_Zpower : zmisc.
+
+Lemma pow_mod_spec : forall n a m, Z_of_N (pow_mod a m n) = a^m mod n.
+Proof with mauto.
+ induction m;simpl;intros...
+ rewrite Zmult_mod; auto with zmisc.
+ rewrite (Zmult_mod (a^m)); auto with zmisc. rewrite <- IHm.
+ destruct (pow_mod a m n);simpl...
+ rewrite Zmult_mod; auto with zmisc.
+ rewrite <- IHm. destruct (pow_mod a m n);simpl...
+Qed.
+Hint Rewrite pow_mod_spec Zpower_0 : zmisc.
+
+Lemma Npow_mod_spec : forall a p n, Z_of_N (Npow_mod a p n) = a^p mod n.
+Proof with mauto.
+ intros a p n;destruct a;simpl ...
+Qed.
+Hint Rewrite Npow_mod_spec : zmisc.
+
+Lemma iter_Npow_mod_spec : forall n q p a,
+ Z_of_N (iter_pos p N (fun x : N => Npow_mod x q n) a) = a^q^p mod n.
+Proof with mauto.
+ induction p;simpl;intros ...
+ repeat rewrite IHp.
+ rewrite (Zpower_mod ((a ^ q ^ p) ^ q ^ p));auto with zmisc.
+ rewrite (Zpower_mod (a ^ q ^ p))...
+ repeat rewrite IHp...
+Qed.
+Hint Rewrite iter_Npow_mod_spec : zmisc.
+
+
+Lemma fold_pow_mod_spec : forall (n:positive) l (a:N),
+ Z_of_N a = a mod n ->
+ Z_of_N (fold_pow_mod a l n) = a^(mkProd' l) mod n.
+Proof with mauto.
+ unfold fold_pow_mod;induction l;simpl;intros ...
+ rewrite IHl...
+Qed.
+Hint Rewrite fold_pow_mod_spec : zmisc.
+
+Lemma pow_mod_pred_spec : forall (n:positive) l (a:N),
+ Z_of_N a = a mod n ->
+ Z_of_N (pow_mod_pred a l n) = a^(mkProd_pred l) mod n.
+Proof with mauto.
+ unfold pow_mod_pred;induction l;simpl;intros ...
+ destruct a as (q,p);simpl.
+ destruct (p ?= 1)%P; rewrite IHl...
+Qed.
+Hint Rewrite pow_mod_pred_spec : zmisc.
+
+Lemma mkProd_pred_mkProd : forall l,
+ (mkProd_pred l)*(mkProd' l) = mkProd l.
+Proof with mauto.
+ induction l;simpl;intros ...
+ generalize (pos_eq_1_spec (snd a)); destruct (snd a ?= 1)%P;intros.
+ rewrite H...
+ replace (mkProd_pred l * (fst a * mkProd' l)) with
+ (fst a *(mkProd_pred l * mkProd' l));try ring.
+ rewrite IHl...
+ rewrite Zmult_assoc. rewrite times_Zmult.
+ rewrite (Zmult_comm (pow (fst a) (Ppred (snd a)) * mkProd_pred l)).
+ rewrite Zmult_assoc. rewrite pow_Zpower. rewrite <-Ppred_Zminus;trivial.
+ rewrite <- Zpower_Zsucc; try omega.
+ replace (Zsucc (snd a - 1)) with ((snd a - 1)+1).
+ replace ((snd a - 1)+1) with (Zpos (snd a)) ...
+ rewrite <- IHl;repeat rewrite Zmult_assoc ...
+ destruct (snd a - 1);trivial.
+ assert (1 < snd a); auto with zarith.
+Qed.
+Hint Rewrite mkProd_pred_mkProd : zmisc.
+
+Lemma lt_Zmod : forall p n, 0 <= p < n -> p mod n = p.
+Proof with mauto.
+ intros a b H.
+ assert ( 0 <= a mod b < b).
+ apply Z_mod_lt...
+ destruct (mod_unique b (a/b) (a mod b) 0 a H0 H)...
+ rewrite <- Z_div_mod_eq...
+Qed.
+
+Opaque Zminus.
+Lemma Npred_mod_spec : forall p n, Z_of_N p < Zpos n ->
+ 1 < Zpos n -> Z_of_N (Npred_mod p n) = (p - 1) mod n.
+Proof with mauto.
+ destruct p;intros;simpl.
+ rewrite <- Ppred_Zminus...
+ change (-1) with (0 -1). rewrite <- (Z_mod_same n) ...
+ pattern 1 at 2;rewrite <- (lt_Zmod 1 n) ...
+ symmetry;apply lt_Zmod.
+Transparent Zminus.
+ omega.
+ assert (H1 := pos_eq_1_spec p);destruct (p?=1)%P.
+ rewrite H1 ...
+ unfold Z_of_N;rewrite <- Ppred_Zminus...
+ simpl in H;symmetry; apply (lt_Zmod (p-1) n)...
+ assert (1 < p); auto with zarith.
+Qed.
+Hint Rewrite Npred_mod_spec : zmisc.
+
+Lemma times_mod_spec : forall x y n, Z_of_N (times_mod x y n) = (x * y) mod n.
+Proof with mauto.
+ intros; destruct x ...
+ destruct y;simpl ...
+Qed.
+Hint Rewrite times_mod_spec : zmisc.
+
+Lemma snd_all_pow_mod :
+ forall n l (prod a :N),
+ a mod (Zpos n) = a ->
+ Z_of_N (snd (all_pow_mod prod a l n)) = (a^(mkProd' l)) mod n.
+Proof with mauto.
+ induction l;simpl;intros...
+ destruct a as (q,p);simpl.
+ rewrite IHl...
+Qed.
+
+Lemma fold_aux : forall a N (n:positive) l prod,
+ fold_left
+ (fun (r : Z) (k : positive * positive) =>
+ r * (a ^(N / fst k) - 1) mod n) l (prod mod n) mod n =
+ fold_left
+ (fun (r : Z) (k : positive * positive) =>
+ r * (a^(N / fst k) - 1)) l prod mod n.
+Proof with mauto.
+ induction l;simpl;intros ...
+Qed.
+
+Lemma fst_all_pow_mod :
+ forall (n a:positive) l (R:positive) (prod A :N),
+ 1 < n ->
+ Z_of_N prod = prod mod n ->
+ Z_of_N A = a^R mod n ->
+ Z_of_N (fst (all_pow_mod prod A l n)) =
+ (fold_left
+ (fun r (k:positive*positive) =>
+ (r * (a ^ (R* mkProd' l / (fst k)) - 1))) l prod) mod n.
+Proof with mauto.
+ induction l;simpl;intros...
+ destruct a0 as (q,p);simpl.
+ assert (Z_of_N A = A mod n).
+ rewrite H1 ...
+ rewrite (IHl (R * q)%positive)...
+ pattern (q * mkProd' l) at 2;rewrite (Zmult_comm q).
+ repeat rewrite Zmult_assoc.
+ rewrite Z_div_mult;auto with zmisc zarith.
+ rewrite <- fold_aux.
+ rewrite <- (fold_aux a (R * q * mkProd' l) n l (prod * (a ^ (R * mkProd' l) - 1)))...
+ assert ( ((prod * (A ^ mkProd' l - 1)) mod n) =
+ ((prod * ((a ^ R) ^ mkProd' l - 1)) mod n)).
+ repeat rewrite (Zmult_mod prod);auto with zmisc.
+ rewrite Zminus_mod;auto with zmisc.
+ rewrite (Zminus_mod ((a ^ R) ^ mkProd' l));auto with zmisc.
+ rewrite (Zpower_mod (a^R));auto with zmisc. rewrite H1...
+ rewrite H3...
+ rewrite H1 ...
+Qed.
+
+
+Lemma is_odd_Zodd : forall p, is_odd p = true -> Zodd p.
+Proof.
+ destruct p;intros;simpl;trivial;discriminate.
+Qed.
+
+Lemma is_even_Zeven : forall p, is_even p = true -> Zeven p.
+Proof.
+ destruct p;intros;simpl;trivial;discriminate.
+Qed.
+
+Lemma lt_square : forall x y, 0 < x -> x < y -> x*x < y*y.
+Proof.
+ intros; apply Zlt_trans with (x*y).
+ apply Zmult_lt_compat_l;trivial.
+ apply Zmult_lt_compat_r;trivial. omega.
+Qed.
+
+Lemma le_square : forall x y, 0 <= x -> x <= y -> x*x <= y*y.
+Proof.
+ intros; apply Zle_trans with (x*y).
+ apply Zmult_le_compat_l;trivial.
+ apply Zmult_le_compat_r;trivial. omega.
+Qed.
+
+Lemma borned_square : forall x y, 0 <= x -> 0 <= y ->
+ x*x < y*y < (x+1)*(x+1) -> False.
+Proof.
+ intros;destruct (Z_lt_ge_dec x y) as [z|z].
+ assert (x + 1 <= y). omega.
+ assert (0 <= x+1). omega.
+ assert (H4 := le_square _ _ H3 H2). omega.
+ assert (H4 := le_square _ _ H0 (Zge_le _ _ z)). omega.
+Qed.
+
+Lemma not_square : forall (sqrt:positive) n, sqrt * sqrt < n < (sqrt+1)*(sqrt + 1) -> ~(isSquare n).
+Proof.
+ intros sqrt n H (y,H0).
+ destruct (Z_lt_ge_dec 0 y).
+ apply (borned_square sqrt y);mauto.
+ assert (y*y = (-y)*(-y)). ring. rewrite H1 in H0;clear H1.
+ apply (borned_square sqrt (-y));mauto.
+Qed.
+
+Ltac spec_dec :=
+ repeat match goal with
+ | [H:(?x ?= ?y)%P = _ |- _] =>
+ generalize (is_eq_spec x y);
+ rewrite H;clear H;simpl; autorewrite with zmisc;
+ intro
+ | [H:(?x ?< ?y)%P = _ |- _] =>
+ generalize (is_lt_spec x y);
+ rewrite H; clear H;simpl; autorewrite with zmisc;
+ intro
+ end.
+
+Ltac elimif :=
+ match goal with
+ | [H: (if ?b then _ else _) = _ |- _] =>
+ let H1 := fresh "H" in
+ (CaseEq b;intros H1; rewrite H1 in H;
+ try discriminate H); elimif
+ | _ => spec_dec
+ end.
+
+Lemma check_s_r_correct : forall s r sqrt, check_s_r s r sqrt = true ->
+ Z_of_N s = 0 \/ ~ isSquare (r * r - 8 * s).
+Proof.
+ unfold check_s_r;intros.
+ destruct s as [|s]; trivial;auto.
+ right;CaseEq (square r - xO (xO (xO s)));[intros H1|intros p1 H1| intros p1 H1];
+ rewrite H1 in H;try discriminate H.
+ elimif.
+ assert (Zpos (xO (xO (xO s))) = 8 * s). repeat rewrite Zpos_xO_add;ring.
+ generalizeclear H1; rewrite H2;mauto;intros.
+ apply (not_square sqrt).
+ rewrite H1;auto.
+ intros (y,Heq).
+ generalize H1 Heq;mauto.
+ unfold Z_of_N.
+ match goal with |- ?x = _ -> ?y = _ -> _ =>
+ replace x with y; try ring
+ end.
+ intros Heq1;rewrite Heq1;intros Heq2.
+ destruct y;discriminate Heq2.
+Qed.
+
+Opaque Zplus Pplus.
+Lemma in_mkProd_prime_div_in :
+ forall p:positive, prime p ->
+ forall (l:dec_prime),
+ (forall k, In k l -> prime (fst k)) ->
+ Zdivide p (mkProd l) -> exists n,In (p, n) l.
+Proof with mauto.
+ induction l;simpl ...
+ intros _ H1; absurd (p <= 1).
+ apply Zlt_not_le; apply Zlt_le_trans with 2; try apply prime_ge_2; auto with zarith.
+ apply Zdivide_le; auto with zarith.
+ intros; case prime_mult with (2 := H1); auto with zarith; intros H2.
+ exists (snd a);left.
+ destruct a;simpl in *.
+ assert (Zpos p = Zpos p0).
+ rewrite (prime_div_Zpower_prime p1 p p0)...
+ apply (H0 (p0,p1));auto.
+ inversion H3...
+ destruct IHl as (n,H3)...
+ exists n...
+Qed.
+
+Lemma gcd_Zis_gcd : forall a b:positive, (Zis_gcd b a (gcd b a)%P).
+Proof with mauto.
+ intros a;assert (Hacc := Zwf_pos a);induction Hacc;rename x into a;intros.
+ generalize (div_eucl_spec b a)...
+ rewrite <- (Pmod_div_eucl b a).
+ CaseEq (b mod a)%P;[intros Heq|intros r Heq]; intros (H1,H2).
+ simpl in H1;rewrite Zplus_0_r in H1.
+ rewrite (gcd_mod0 _ _ Heq).
+ constructor;mauto.
+ apply Zdivide_intro with (fst (b/a)%P);trivial.
+ rewrite (gcd_mod _ _ _ Heq).
+ rewrite H1;apply Zis_gcd_sym.
+ rewrite Zmult_comm;apply Zis_gcd_for_euclid2;simpl in *.
+ apply Zis_gcd_sym;auto.
+Qed.
+
+Lemma test_pock_correct : forall N a dec sqrt,
+ (forall k, In k dec -> prime (Zpos (fst k))) ->
+ test_pock N a dec sqrt = true ->
+ prime N.
+Proof with mauto.
+ unfold test_pock;intros.
+ elimif.
+ generalize (div_eucl_spec (Ppred N) (mkProd dec));
+ destruct ((Ppred N) / (mkProd dec))%P as (R1,n);simpl;mauto;intros (H2,H3).
+ destruct R1 as [|R1];try discriminate H0.
+ destruct n;try discriminate H0.
+ elimif.
+ generalize (div_eucl_spec R1 (xO (mkProd dec)));
+ destruct ((R1 / xO (mkProd dec))%P) as (s,r');simpl;mauto;intros (H7,H8).
+ destruct r' as [|r];try discriminate H0.
+ generalize (fst_all_pow_mod N a dec (R1*mkProd_pred dec) 1
+ (pow_mod_pred (pow_mod a R1 N) dec N)).
+ generalize (snd_all_pow_mod N dec 1 (pow_mod_pred (pow_mod a R1 N) dec N)).
+ destruct (all_pow_mod 1 (pow_mod_pred (pow_mod a R1 N) dec N) dec N) as
+ (prod,aNm1);simpl...
+ destruct prod as [|prod];try discriminate H0.
+ destruct aNm1 as [|aNm1];try discriminate H0;elimif.
+ simpl in H2;rewrite Zplus_0_r in H2.
+ rewrite <- Ppred_Zminus in H2;try omega.
+ rewrite <- Zmult_assoc;rewrite mkProd_pred_mkProd.
+ intros H12;assert (a^(N-1) mod N = 1).
+ pattern 1 at 2;rewrite <- H9;symmetry.
+ rewrite H2;rewrite H12 ...
+ rewrite <- Zpower_mult...
+ clear H12.
+ intros H14.
+ match type of H14 with _ -> _ -> _ -> ?X =>
+ assert (H12:X); try apply H14; clear H14
+ end...
+ rewrite Zmod_small...
+ assert (1 < mkProd dec).
+ assert (H14 := Zlt_0_pos (mkProd dec)).
+ assert (1 <= mkProd dec)...
+ destruct (Zle_lt_or_eq _ _ H15)...
+ inversion H16. rewrite <- H18 in H5;discriminate H5.
+ simpl in H8.
+ assert (Z_of_N s = R1 / (2 * mkProd dec) /\ Zpos r = R1 mod (2 * mkProd dec)).
+ apply mod_unique with (2 * mkProd dec);auto with zarith.
+ apply Z_mod_lt ...
+ rewrite <- Z_div_mod_eq... rewrite H7. simpl;ring.
+ destruct H15 as (H15,Heqr).
+ apply PocklingtonExtra with (F1:=mkProd dec) (R1:=R1) (m:=1);
+ auto with zmisc zarith.
+ rewrite H2;ring.
+ apply is_even_Zeven...
+ apply is_odd_Zodd...
+ intros p; case p; clear p.
+ intros HH; contradict HH.
+ apply not_prime_0.
+ 2: intros p (V1, _); contradict V1; apply Zle_not_lt; red; simpl; intros;
+ discriminate.
+ intros p Hprime Hdec; exists (Zpos a);repeat split; auto with zarith.
+ apply Zis_gcd_gcd; auto with zarith.
+ change (rel_prime (a ^ ((N - 1) / p) - 1) N).
+ match type of H12 with _ = ?X mod _ =>
+ apply rel_prime_div with (p := X); auto with zarith
+ end.
+ apply rel_prime_mod_rev; auto with zarith.
+ red.
+ pattern 1 at 3; rewrite <- H10; rewrite <- H12.
+ apply Pmod.gcd_Zis_gcd.
+ destruct (in_mkProd_prime_div_in _ Hprime _ H Hdec) as (q,Hin).
+ rewrite <- H2.
+ match goal with |- context [fold_left ?f _ _] =>
+ apply (ListAux.fold_left_invol_in _ _ f (fun k => Zdivide (a ^ ((N - 1) / p) - 1) k))
+ with (b := (p, q)); auto with zarith
+ end.
+ rewrite <- Heqr.
+ generalizeclear H0; ring_simplify
+ (((mkProd dec + mkProd dec + r + 1) * mkProd dec + r) * mkProd dec + 1)
+ ((1 * mkProd dec + 1) * (2 * mkProd dec * mkProd dec + (r - 1) * mkProd dec + 1))...
+ rewrite <- H15;rewrite <- Heqr.
+ apply check_s_r_correct with sqrt ...
+Qed.
+
+Lemma is_in_In :
+ forall p lc, is_in p lc = true -> exists c, In c lc /\ p = nprim c.
+Proof.
+ induction lc;simpl;try (intros;discriminate).
+ intros;elimif.
+ exists a;split;auto. inversion H0;trivial.
+ destruct (IHlc H) as [c [H1 H2]];exists c;auto.
+Qed.
+
+Lemma all_in_In :
+ forall lc lp, all_in lc lp = true ->
+ forall pq, In pq lp -> exists c, In c lc /\ fst pq = nprim c.
+Proof.
+ induction lp;simpl. intros H pq HF;elim HF.
+ intros;destruct a;elimif.
+ destruct H0;auto.
+ rewrite <- H0;simpl;apply is_in_In;trivial.
+Qed.
+
+Lemma test_Certif_In_Prime :
+ forall lc, test_Certif lc = true ->
+ forall c, In c lc -> prime (nprim c).
+Proof with mauto.
+ induction lc;simpl;intros. elim H0.
+ destruct H0.
+ subst c;destruct a;simpl...
+ elimif.
+ CaseEq (Mp p);[intros Heq|intros N' Heq|intros N' Heq];rewrite Heq in H;
+ try discriminate H. elimif.
+ CaseEq (SS p);[intros Heq'|intros N'' Heq'|intros N'' Heq'];rewrite Heq' in H;
+ try discriminate H.
+ rewrite H2;rewrite <- Heq.
+apply LucasLehmer;trivial.
+(destruct p; try discriminate H1).
+simpl in H1; generalize (is_lt_spec 2 p); rewrite H1; auto.
+elimif.
+apply (test_pock_correct N a d p); mauto.
+ intros k Hin;destruct (all_in_In _ _ H1 _ Hin) as (c,(H2,H3)).
+ rewrite H3;auto.
+discriminate.
+discriminate.
+ destruct a;elimif;auto.
+discriminate.
+discriminate.
+Qed.
+
+Lemma Pocklington_refl :
+ forall c lc, test_Certif (c::lc) = true -> prime (nprim c).
+Proof.
+ intros c lc Heq;apply test_Certif_In_Prime with (c::lc);trivial;simpl;auto.
+Qed.
+
diff --git a/coqprime/PrimalityTest/Proth.v b/coqprime/PrimalityTest/Proth.v
new file mode 100644
index 000000000..b087f1854
--- /dev/null
+++ b/coqprime/PrimalityTest/Proth.v
@@ -0,0 +1,120 @@
+
+(*************************************************************)
+(* This file is distributed under the terms of the *)
+(* GNU Lesser General Public License Version 2.1 *)
+(*************************************************************)
+(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
+(*************************************************************)
+
+(**********************************************************************
+ Proth.v
+
+ Proth's Test
+
+ Definition: ProthTest
+ **********************************************************************)
+Require Import ZArith.
+Require Import ZCAux.
+Require Import Pocklington.
+
+Open Scope Z_scope.
+
+Theorem ProthTest: forall h k a, let n := h * 2 ^ k + 1 in 1 < a -> 0 < h < 2 ^k -> (a ^ ((n - 1) / 2) + 1) mod n = 0 -> prime n.
+intros h k a n; unfold n; intros H H1 H2.
+assert (Hu: 0 < h * 2 ^ k).
+apply Zmult_lt_O_compat; auto with zarith.
+assert (Hu1: 0 < k).
+case (Zle_or_lt k 0); intros Hv; auto.
+generalize H1 Hv; case k; simpl.
+intros (Hv1, Hv2); contradict Hv2; auto with zarith.
+intros p1 _ Hv1; contradict Hv1; auto with zarith.
+intros p (Hv1, Hv2); contradict Hv2; auto with zarith.
+apply PocklingtonCorollary1 with (F1 := 2 ^ k) (R1 := h); auto with zarith.
+ring.
+apply Zlt_le_trans with ((h + 1) * 2 ^ k); auto with zarith.
+rewrite Zmult_plus_distr_l; apply Zplus_lt_compat_l.
+rewrite Zmult_1_l; apply Zlt_le_trans with 2; auto with zarith.
+intros p H3 H4.
+generalize H2; replace (h * 2 ^ k + 1 - 1) with (h * 2 ^k); auto with zarith; clear H2; intros H2.
+exists a; split; auto; split.
+pattern (h * 2 ^k) at 1; rewrite (Zdivide_Zdiv_eq 2 (h * 2 ^ k)); auto with zarith.
+rewrite (Zmult_comm 2); rewrite Zpower_mult; auto with zarith.
+rewrite Zpower_mod; auto with zarith.
+assert (tmp: forall p, p = (p + 1) -1); auto with zarith; rewrite (fun x => (tmp (a ^ x))).
+rewrite Zminus_mod; auto with zarith.
+rewrite H2.
+rewrite (Zmod_small 1); auto with zarith.
+rewrite <- Zpower_mod; auto with zarith.
+rewrite Zmod_small; auto with zarith.
+simpl; unfold Zpower_pos; simpl; auto with zarith.
+apply Z_div_pos; auto with zarith.
+apply Zdivide_trans with (2 ^ k).
+apply Zpower_divide; auto with zarith.
+apply Zdivide_factor_l; auto with zarith.
+apply Zis_gcd_gcd; auto with zarith.
+apply Zis_gcd_intro; auto with zarith.
+intros x HD1 HD2.
+assert (Hd1: p = 2).
+apply prime_div_Zpower_prime with (4 := H4); auto with zarith.
+apply prime_2.
+assert (Hd2: (x | 2)).
+replace 2 with ((a ^ (h * 2 ^ k / 2) + 1) - (a ^ (h * 2 ^ k/ 2) - 1)); auto with zarith.
+apply Zdivide_minus_l; auto.
+apply Zdivide_trans with (1 := HD2).
+apply Zmod_divide; auto with zarith.
+pattern 2 at 2; rewrite <- Hd1; auto.
+replace 1 with ((h * 2 ^k + 1) - (h * 2 ^ k)); auto with zarith.
+apply Zdivide_minus_l; auto.
+apply Zdivide_trans with (1 := Hd2); auto.
+apply Zdivide_trans with (2 ^ k).
+apply Zpower_divide; auto with zarith.
+apply Zdivide_factor_l; auto with zarith.
+Qed.
+
+
+Definition proth_test h k a :=
+ let n := h * 2 ^ k + 1 in
+ if (Z_lt_dec 1 a) then
+ if (Z_lt_dec 0 h) then
+ if (Z_lt_dec h (2 ^k)) then
+ if Z_eq_dec (Zpow_mod a ((n - 1) / 2) n) (n - 1) then true
+ else false else false else false else false.
+
+
+Theorem ProthTestOp: forall h k a, proth_test h k a = true -> prime (h * 2 ^ k + 1).
+intros h k a; unfold proth_test.
+repeat match goal with |- context[if ?X then _ else _] => case X end; try (intros; discriminate).
+intros H1 H2 H3 H4 _.
+assert (Hu: 0 < h * 2 ^ k).
+apply Zmult_lt_O_compat; auto with zarith.
+apply ProthTest with (a := a); auto.
+rewrite Zplus_mod; auto with zarith.
+rewrite <- Zpow_mod_Zpower_correct; auto with zarith.
+rewrite H1.
+rewrite (Zmod_small 1); auto with zarith.
+replace (h * 2 ^ k + 1 - 1 + 1) with (h * 2 ^ k + 1); auto with zarith.
+apply Zdivide_mod; auto with zarith.
+apply Z_div_pos; auto with zarith.
+Qed.
+
+Theorem prime5: prime 5.
+exact (ProthTestOp 1 2 2 (refl_equal _)).
+Qed.
+
+Theorem prime17: prime 17.
+exact (ProthTestOp 1 4 3 (refl_equal _)).
+Qed.
+
+Theorem prime257: prime 257.
+exact (ProthTestOp 1 8 3 (refl_equal _)).
+Qed.
+
+Theorem prime65537: prime 65537.
+exact (ProthTestOp 1 16 3 (refl_equal _)).
+Qed.
+
+(* Too tough !!
+Theorem prime4294967297: prime 4294967297.
+exact (ProthTestOp 1 32 3 (refl_equal _)).
+Qed.
+*)
diff --git a/coqprime/PrimalityTest/Root.v b/coqprime/PrimalityTest/Root.v
new file mode 100644
index 000000000..321865ba1
--- /dev/null
+++ b/coqprime/PrimalityTest/Root.v
@@ -0,0 +1,239 @@
+
+(*************************************************************)
+(* This file is distributed under the terms of the *)
+(* GNU Lesser General Public License Version 2.1 *)
+(*************************************************************)
+(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
+(*************************************************************)
+
+(***********************************************************************
+ Root.v
+
+ Proof that a polynomial has at most n roots
+************************************************************************)
+Require Import ZArith.
+Require Import List.
+Require Import UList.
+Require Import Tactic.
+Require Import Permutation.
+
+Open Scope Z_scope.
+
+Section Root.
+
+Variable A: Set.
+Variable P: A -> Prop.
+Variable plus mult: A -> A -> A.
+Variable op: A -> A.
+Variable zero one: A.
+
+
+Let pol := list A.
+
+Definition toA z :=
+match z with
+ Z0 => zero
+| Zpos p => iter_pos p _ (plus one) zero
+| Zneg p => op (iter_pos p _ (plus one) zero)
+end.
+
+Fixpoint eval (p: pol) (x: A) {struct p} : A :=
+match p with
+ nil => zero
+| a::p1 => plus a (mult x (eval p1 x))
+end.
+
+Fixpoint div (p: pol) (x: A) {struct p} : pol * A :=
+match p with
+ nil => (nil, zero)
+| a::nil => (nil, a)
+| a::p1 =>
+ (snd (div p1 x)::fst (div p1 x),
+ (plus a (mult x (snd (div p1 x)))))
+end.
+
+Hypothesis Pzero: P zero.
+Hypothesis Pplus: forall x y, P x -> P y -> P (plus x y).
+Hypothesis Pmult: forall x y, P x -> P y -> P (mult x y).
+Hypothesis Pop: forall x, P x -> P (op x).
+Hypothesis plus_zero: forall a, P a -> plus zero a = a.
+Hypothesis plus_comm: forall a b, P a -> P b -> plus a b = plus b a.
+Hypothesis plus_assoc: forall a b c, P a -> P b -> P c -> plus a (plus b c) = plus (plus a b) c.
+Hypothesis mult_zero: forall a, P a -> mult zero a = zero.
+Hypothesis mult_comm: forall a b, P a -> P b -> mult a b = mult b a.
+Hypothesis mult_assoc: forall a b c, P a -> P b -> P c -> mult a (mult b c) = mult (mult a b) c.
+Hypothesis mult_plus_distr: forall a b c, P a -> P b -> P c -> mult a (plus b c) = plus (mult a b) (mult a c).
+Hypothesis plus_op_zero: forall a, P a -> plus a (op a) = zero.
+Hypothesis mult_integral: forall a b, P a -> P b -> mult a b = zero -> a = zero \/ b = zero.
+(* Not necessary in Set just handy *)
+Hypothesis A_dec: forall a b: A, {a = b} + {a <> b}.
+
+Theorem eval_P: forall p a, P a -> (forall i, In i p -> P i) -> P (eval p a).
+intros p a Pa; elim p; simpl; auto with datatypes.
+intros a1 l1 Rec H; apply Pplus; auto.
+Qed.
+
+Hint Resolve eval_P.
+
+Theorem div_P: forall p a, P a -> (forall i, In i p -> P i) -> (forall i, In i (fst (div p a)) -> P i) /\ P (snd (div p a)).
+intros p a Pa; elim p; auto with datatypes.
+intros a1 l1; case l1.
+simpl; intuition.
+intros a2 p2 Rec Hi; split.
+case Rec; auto with datatypes.
+intros H H1 i.
+replace (In i (fst (div (a1 :: a2 :: p2) a))) with
+ (snd (div (a2::p2) a) = i \/ In i (fst (div (a2::p2) a))); auto.
+intros [Hi1 | Hi1]; auto.
+rewrite <- Hi1; auto.
+change ( P (plus a1 (mult a (snd (div (a2::p2) a))))); auto with datatypes.
+apply Pplus; auto with datatypes.
+apply Pmult; auto with datatypes.
+case Rec; auto with datatypes.
+Qed.
+
+
+Theorem div_correct:
+ forall p x y, P x -> P y -> (forall i, In i p -> P i) -> eval p y = plus (mult (eval (fst (div p x)) y) (plus y (op x))) (snd (div p x)).
+intros p x y; elim p; simpl.
+intros; rewrite mult_zero; try rewrite plus_zero; auto.
+intros a l; case l; simpl; auto.
+intros _ px py pa; rewrite (fun x => mult_comm x zero); repeat rewrite mult_zero; try apply plus_comm; auto.
+intros a1 l1.
+generalize (div_P (a1::l1) x); simpl.
+match goal with |- context[fst ?A] => case A end; simpl.
+intros q r Hd Rec px py pi.
+assert (pr: P r).
+case Hd; auto.
+assert (pa1: P a1).
+case Hd; auto.
+assert (pey: P (eval q y)).
+apply eval_P; auto.
+case Hd; auto.
+rewrite Rec; auto with datatypes.
+rewrite (fun x y => plus_comm x (plus a y)); try rewrite <- plus_assoc; auto.
+apply f_equal2 with (f := plus); auto.
+repeat rewrite mult_plus_distr; auto.
+repeat (rewrite (fun x y => (mult_comm (plus x y))) || rewrite mult_plus_distr); auto.
+rewrite (fun x => (plus_comm x (mult y r))); auto.
+repeat rewrite plus_assoc; try apply f_equal2 with (f := plus); auto.
+2: repeat rewrite mult_assoc; try rewrite (fun y => mult_comm y (op x));
+ repeat rewrite mult_assoc; auto.
+rewrite (fun z => (plus_comm z (mult (op x) r))); auto.
+repeat rewrite plus_assoc; try apply f_equal2 with (f := plus); auto.
+2: apply f_equal2 with (f := mult); auto.
+repeat rewrite (fun x => mult_comm x r); try rewrite <- mult_plus_distr; auto.
+rewrite (plus_comm (op x)); try rewrite plus_op_zero; auto.
+rewrite (fun x => mult_comm x zero); try rewrite mult_zero; try rewrite plus_zero; auto.
+Qed.
+
+Theorem div_correct_factor:
+ forall p a, (forall i, In i p -> P i) -> P a ->
+ eval p a = zero -> forall x, P x -> eval p x = (mult (eval (fst (div p a)) x) (plus x (op a))).
+intros p a Hp Ha H x px.
+case (div_P p a); auto; intros Hd1 Hd2.
+rewrite (div_correct p a x); auto.
+generalize (div_correct p a a).
+rewrite plus_op_zero; try rewrite (fun x => mult_comm x zero); try rewrite mult_zero; try rewrite plus_zero; try rewrite H; auto.
+intros H1; rewrite <- H1; auto.
+rewrite (fun x => plus_comm x zero); auto.
+Qed.
+
+Theorem length_decrease: forall p x, p <> nil -> (length (fst (div p x)) < length p)%nat.
+intros p x; elim p; simpl; auto.
+intros H1; case H1; auto.
+intros a l; case l; simpl; auto.
+intros a1 l1.
+match goal with |- context[fst ?A] => case A end; simpl; auto with zarith.
+intros p1 _ H H1.
+apply lt_n_S; apply H; intros; discriminate.
+Qed.
+
+Theorem root_max:
+forall p l, ulist l -> (forall i, In i p -> P i) -> (forall i, In i l -> P i) ->
+ (forall x, In x l -> eval p x = zero) -> (length p <= length l)%nat -> forall x, P x -> eval p x = zero.
+intros p l; generalize p; elim l; clear l p; simpl; auto.
+intros p; case p; simpl; auto.
+intros a p1 _ _ _ _ H; contradict H; auto with arith.
+intros a p1 Rec p; case p.
+simpl; auto.
+intros a1 p2 H H1 H2 H3 H4 x px.
+assert (Hu: eval (a1 :: p2) a = zero); auto with datatypes.
+rewrite (div_correct_factor (a1 :: p2) a); auto with datatypes.
+match goal with |- mult ?X _ = _ => replace X with zero end; try apply mult_zero; auto.
+apply sym_equal; apply Rec; auto with datatypes.
+apply ulist_inv with (1 := H).
+intros i Hi; case (div_P (a1 :: p2) a); auto.
+intros x1 H5; case (mult_integral (eval (fst (div (a1 :: p2) a)) x1) (plus x1 (op a))); auto.
+apply eval_P; auto.
+intros i Hi; case (div_P (a1 :: p2) a); auto.
+rewrite <- div_correct_factor; auto.
+intros H6; case (ulist_app_inv _ (a::nil) p1 x1); simpl; auto.
+left.
+apply trans_equal with (plus zero x1); auto.
+rewrite <- (plus_op_zero a); try rewrite <- plus_assoc; auto.
+rewrite (fun x => plus_comm (op x)); try rewrite H6; try rewrite plus_comm; auto.
+apply sym_equal; apply plus_zero; auto.
+apply lt_n_Sm_le;apply lt_le_trans with (length (a1 :: p2)); auto with zarith.
+apply length_decrease; auto with datatypes.
+Qed.
+
+Theorem root_max_is_zero:
+forall p l, ulist l -> (forall i, In i p -> P i) -> (forall i, In i l -> P i) ->
+ (forall x, In x l -> eval p x = zero) -> (length p <= length l)%nat -> forall x, (In x p) -> x = zero.
+intros p l; generalize p; elim l; clear l p; simpl; auto.
+intros p; case p; simpl; auto.
+intros _ _ _ _ _ x H; case H.
+intros a p1 _ _ _ _ H; contradict H; auto with arith.
+intros a p1 Rec p; case p.
+simpl; auto.
+intros _ _ _ _ _ x H; case H.
+simpl; intros a1 p2 H H1 H2 H3 H4 x H5.
+assert (Ha1: a1 = zero).
+assert (Hu: (eval (a1::p2) zero = zero)).
+apply root_max with (l := a :: p1); auto.
+rewrite <- Hu; simpl; rewrite mult_zero; try rewrite plus_comm; sauto.
+case H5; clear H5; intros H5; subst; auto.
+apply Rec with p2; auto with arith.
+apply ulist_inv with (1 := H).
+intros x1 Hx1.
+case (In_dec A_dec zero p1); intros Hz.
+case (in_permutation_ex _ zero p1); auto; intros p3 Hp3.
+apply root_max with (l := a::p3); auto.
+apply ulist_inv with zero.
+apply ulist_perm with (a::p1); auto.
+apply permutation_trans with (a:: (zero:: p3)); auto.
+apply permutation_skip; auto.
+apply permutation_sym; auto.
+simpl; intros x2 [Hx2 | Hx2]; subst; auto.
+apply H2; right; apply permutation_in with (1 := Hp3); auto with datatypes.
+simpl; intros x2 [Hx2 | Hx2]; subst.
+case (mult_integral x2 (eval p2 x2)); auto.
+rewrite <- H3 with x2; sauto.
+rewrite plus_zero; auto.
+intros H6; case (ulist_app_inv _ (x2::nil) p1 x2) ; auto with datatypes.
+rewrite H6; apply permutation_in with (1 := Hp3); auto with datatypes.
+case (mult_integral x2 (eval p2 x2)); auto.
+apply H2; right; apply permutation_in with (1 := Hp3); auto with datatypes.
+apply eval_P; auto.
+apply H2; right; apply permutation_in with (1 := Hp3); auto with datatypes.
+rewrite <- H3 with x2; sauto; try right.
+apply sym_equal; apply plus_zero; auto.
+apply Pmult; auto.
+apply H2; right; apply permutation_in with (1 := Hp3); auto with datatypes.
+apply eval_P; auto.
+apply H2; right; apply permutation_in with (1 := Hp3); auto with datatypes.
+apply permutation_in with (1 := Hp3); auto with datatypes.
+intros H6; case (ulist_app_inv _ (zero::nil) p3 x2) ; auto with datatypes.
+simpl; apply ulist_perm with (1:= (permutation_sym _ _ _ Hp3)).
+apply ulist_inv with (1 := H).
+rewrite H6; auto with datatypes.
+replace (length (a :: p3)) with (length (zero::p3)); auto.
+rewrite permutation_length with (1 := Hp3); auto with arith.
+case (mult_integral x1 (eval p2 x1)); auto.
+rewrite <- H3 with x1; sauto; try right.
+apply sym_equal; apply plus_zero; auto.
+intros HH; case Hz; rewrite <- HH; auto.
+Qed.
+
+End Root. \ No newline at end of file
diff --git a/coqprime/PrimalityTest/Zp.v b/coqprime/PrimalityTest/Zp.v
new file mode 100644
index 000000000..1e5295191
--- /dev/null
+++ b/coqprime/PrimalityTest/Zp.v
@@ -0,0 +1,411 @@
+
+(*************************************************************)
+(* This file is distributed under the terms of the *)
+(* GNU Lesser General Public License Version 2.1 *)
+(*************************************************************)
+(* Benjamin.Gregoire@inria.fr Laurent.Thery@inria.fr *)
+(*************************************************************)
+
+(**********************************************************************
+ Zp.v
+
+ Build the group of the inversible element on {1, 2, .., n-1}
+ for the multiplication modulo n
+
+ Definition: ZpGroup
+ **********************************************************************)
+Require Import ZArith Znumtheory Zpow_facts.
+Require Import Tactic.
+Require Import Wf_nat.
+Require Import UList.
+Require Import FGroup.
+Require Import EGroup.
+Require Import IGroup.
+Require Import Cyclic.
+Require Import Euler.
+Require Import ZProgression.
+
+Open Scope Z_scope.
+
+Section Zp.
+
+Variable n: Z.
+
+Hypothesis n_pos: 1 < n.
+
+
+(**************************************
+ mkZp m creates {m, m - 1, ..., 0}
+ **************************************)
+
+Fixpoint mkZp_aux (m: nat): list Z:=
+ Z_of_nat m :: match m with O => nil | (S m1) => mkZp_aux m1 end.
+
+(**************************************
+ Some properties of mkZp_aux
+ **************************************)
+
+Theorem mkZp_aux_length: forall m, length (mkZp_aux m) = (m + 1)%nat.
+intros m; elim m; simpl; auto.
+Qed.
+
+Theorem mkZp_aux_in: forall m p, 0 <= p <= Z_of_nat m -> In p (mkZp_aux m).
+intros m; elim m.
+simpl; auto with zarith.
+intros n1 Rec p (H1, H2); case Zle_lt_or_eq with (1 := H2); clear H2; intro H2.
+rewrite inj_S in H2.
+simpl; right; apply Rec; split; auto with zarith.
+rewrite H2; simpl; auto.
+Qed.
+
+Theorem in_mkZp_aux: forall m p, In p (mkZp_aux m) -> 0 <= p <= Z_of_nat m.
+intros m; elim m; clear m.
+simpl; intros p H1; case H1; clear H1; intros H1; subst; auto with zarith.
+intros m1; generalize (inj_S m1); simpl.
+intros H Rec p [H1 | H1].
+rewrite <- H1; rewrite H; auto with zarith.
+rewrite H; case (Rec p); auto with zarith.
+Qed.
+
+Theorem mkZp_aux_ulist: forall m, ulist (mkZp_aux m).
+intros m; elim m; simpl; auto.
+intros m1 H; apply ulist_cons; auto.
+change (~ In (Z_of_nat (S m1)) (mkZp_aux m1)).
+rewrite inj_S; intros H1.
+case in_mkZp_aux with (1 := H1); auto with zarith.
+Qed.
+
+(**************************************
+ mkZp creates {n - 1, ..., 1, 0}
+ **************************************)
+
+Definition mkZp := mkZp_aux (Zabs_nat (n - 1)).
+
+(**************************************
+ Some properties of mkZp
+ **************************************)
+
+Theorem mkZp_length: length mkZp = Zabs_nat n.
+unfold mkZp; rewrite mkZp_aux_length.
+apply inj_eq_rev.
+rewrite inj_plus.
+simpl; repeat rewrite inj_Zabs_nat; auto with zarith.
+repeat rewrite Zabs_eq; auto with zarith.
+Qed.
+
+Theorem mkZp_in: forall p, 0 <= p < n -> In p mkZp.
+intros p (H1, H2); unfold mkZp; apply mkZp_aux_in.
+rewrite inj_Zabs_nat; auto with zarith.
+repeat rewrite Zabs_eq; auto with zarith.
+Qed.
+
+Theorem in_mkZp: forall p, In p mkZp -> 0 <= p < n.
+intros p H; case (in_mkZp_aux (Zabs_nat (n - 1)) p); auto with zarith.
+rewrite inj_Zabs_nat; auto with zarith.
+repeat rewrite Zabs_eq; auto with zarith.
+Qed.
+
+Theorem mkZp_ulist: ulist mkZp.
+unfold mkZp; apply mkZp_aux_ulist; auto.
+Qed.
+
+(**************************************
+ Multiplication of two pairs
+ **************************************)
+
+Definition pmult (p q: Z) := (p * q) mod n.
+
+(**************************************
+ Properties of multiplication
+ **************************************)
+
+Theorem pmult_assoc: forall p q r, (pmult p (pmult q r)) = (pmult (pmult p q) r).
+assert (Hu: 0 < n); try apply Zlt_trans with 1; auto with zarith.
+generalize Zmod_mod; intros H.
+intros p q r; unfold pmult.
+rewrite (Zmult_mod p); auto.
+repeat rewrite Zmod_mod; auto.
+rewrite (Zmult_mod q); auto.
+rewrite <- Zmult_mod; auto.
+rewrite Zmult_assoc.
+rewrite (Zmult_mod (p * (q mod n))); auto.
+rewrite (Zmult_mod ((p * q) mod n)); auto.
+eq_tac; auto.
+eq_tac; auto.
+rewrite (Zmult_mod p); sauto.
+rewrite Zmod_mod; auto.
+rewrite <- Zmult_mod; sauto.
+Qed.
+
+Theorem pmult_1_l: forall p, In p mkZp -> pmult 1 p = p.
+intros p H; unfold pmult; rewrite Zmult_1_l.
+apply Zmod_small.
+case (in_mkZp p); auto with zarith.
+Qed.
+
+Theorem pmult_1_r: forall p, In p mkZp -> pmult p 1 = p.
+intros p H; unfold pmult; rewrite Zmult_1_r.
+apply Zmod_small.
+case (in_mkZp p); auto with zarith.
+Qed.
+
+Theorem pmult_comm: forall p q, pmult p q = pmult q p.
+intros p q; unfold pmult; rewrite Zmult_comm; auto.
+Qed.
+
+Definition Lrel := isupport_aux _ pmult mkZp 1 Z_eq_dec (progression Zsucc 0 (Zabs_nat n)).
+
+Theorem rel_prime_is_inv:
+ forall a, is_inv Z pmult mkZp 1 Z_eq_dec a = if (rel_prime_dec a n) then true else false.
+assert (Hu: 0 < n); try apply Zlt_trans with 1; auto with zarith.
+intros a; case (rel_prime_dec a n); intros H.
+assert (H1: Bezout a n 1); try apply rel_prime_bezout; auto.
+inversion H1 as [c d Hcd]; clear H1.
+assert (pmult (c mod n) a = 1).
+unfold pmult; rewrite Zmult_mod; try rewrite Zmod_mod; auto.
+rewrite <- Zmult_mod; auto.
+replace (c * a) with (1 + (-d) * n).
+rewrite Z_mod_plus; auto with zarith.
+rewrite Zmod_small; auto with zarith.
+rewrite <- Hcd; ring.
+apply is_inv_true with (a := (c mod n)); auto.
+apply mkZp_in; auto with zarith.
+exact pmult_1_l.
+exact pmult_1_r.
+rewrite pmult_comm; auto.
+apply mkZp_in; auto with zarith.
+apply Z_mod_lt; auto with zarith.
+apply is_inv_false.
+intros c H1; left; intros H2; contradict H.
+apply bezout_rel_prime.
+apply Bezout_intro with c (- (Zdiv (c * a) n)).
+pattern (c * a) at 1; rewrite (Z_div_mod_eq (c * a) n); auto with zarith.
+unfold pmult in H2; rewrite (Zmult_comm c); try rewrite H2.
+ring.
+Qed.
+
+(**************************************
+ We are now ready to build our group
+ **************************************)
+
+Definition ZPGroup : (FGroup pmult).
+apply IGroup with (support := mkZp) (e:= 1).
+exact Z_eq_dec.
+apply mkZp_ulist.
+apply mkZp_in; auto with zarith.
+intros a b H1 H2; apply mkZp_in.
+unfold pmult; apply Z_mod_lt; auto with zarith.
+intros; apply pmult_assoc.
+exact pmult_1_l.
+exact pmult_1_r.
+Defined.
+
+Theorem in_ZPGroup: forall p, rel_prime p n -> 0 <= p < n -> In p ZPGroup.(s).
+intros p H (H1, H2); unfold ZPGroup; simpl.
+apply isupport_is_in.
+generalize (rel_prime_is_inv p); case (rel_prime_dec p); auto.
+apply mkZp_in; auto with zarith.
+Qed.
+
+
+Theorem phi_is_length: phi n = Z_of_nat (length Lrel).
+assert (Hu: 0 < n); try apply Zlt_trans with 1; auto with zarith.
+rewrite phi_def_with_0; auto.
+unfold Zsum, Lrel; rewrite Zle_imp_le_bool; auto with zarith.
+replace (1 + (n - 1) - 0) with n; auto with zarith.
+elim (progression Zsucc 0 (Zabs_nat n)); simpl; auto.
+intros a l1 Rec.
+rewrite Rec.
+rewrite rel_prime_is_inv.
+case (rel_prime_dec a n); auto with zarith.
+simpl length; rewrite inj_S; auto with zarith.
+Qed.
+
+Theorem phi_is_order: phi n = g_order ZPGroup.
+unfold g_order; rewrite phi_is_length.
+eq_tac; apply permutation_length.
+apply ulist_incl2_permutation.
+unfold Lrel; apply isupport_aux_ulist.
+apply ulist_Zprogression; auto.
+apply ZPGroup.(unique_s).
+intros a H; unfold ZPGroup; simpl.
+apply isupport_is_in.
+unfold Lrel in H; apply isupport_aux_is_inv_true with (1 := H).
+apply mkZp_in; auto.
+assert (In a (progression Zsucc 0 (Zabs_nat n))).
+apply (isupport_aux_incl _ pmult mkZp 1 Z_eq_dec); auto.
+split.
+apply Zprogression_le_init with (1 := H0).
+replace n with (0 + Z_of_nat (Zabs_nat n)).
+apply Zprogression_le_end with (1 := H0).
+rewrite inj_Zabs_nat; auto with zarith.
+rewrite Zabs_eq; auto with zarith.
+intros a H; unfold Lrel; simpl.
+apply isupport_aux_is_in.
+simpl in H; apply isupport_is_inv_true with (1 := H).
+apply in_Zprogression.
+rewrite Zplus_0_l; rewrite inj_Zabs_nat; auto with zarith.
+rewrite Zabs_eq; auto with zarith.
+assert (In a mkZp).
+apply (isupport_aux_incl _ pmult mkZp 1 Z_eq_dec); auto.
+apply in_mkZp; auto.
+Qed.
+
+Theorem Zp_cyclic: prime n -> cyclic Z_eq_dec ZPGroup.
+intros H1.
+unfold ZPGroup, pmult;
+generalize (cyclic_field _ (fun x y => (x + y) mod n) (fun x y => (x * y) mod n) (fun x => (-x) mod n) 0);
+unfold IA; intros tmp; apply tmp; clear tmp; auto.
+intros; discriminate.
+apply mkZp_in; auto with zarith.
+intros; apply mkZp_in; auto with zarith.
+apply Z_mod_lt; auto with zarith.
+intros; rewrite Zplus_0_l; auto.
+apply Zmod_small; auto.
+apply in_mkZp; auto.
+intros; rewrite Zplus_comm; auto.
+intros a b c Ha Hb Hc.
+pattern a at 1; rewrite <- (Zmod_small a n); auto with zarith.
+pattern c at 2; rewrite <- (Zmod_small c n); auto with zarith.
+repeat rewrite <- Zplus_mod; auto with zarith.
+eq_tac; auto with zarith.
+apply in_mkZp; auto.
+apply in_mkZp; auto.
+intros; eq_tac; auto with zarith.
+intros a b c Ha Hb Hc.
+pattern a at 1; rewrite <- (Zmod_small a n); auto with zarith.
+repeat rewrite <- Zmult_mod; auto with zarith.
+repeat rewrite <- Zplus_mod; auto with zarith.
+eq_tac; auto with zarith.
+apply in_mkZp; auto.
+intros; apply mkZp_in; apply Z_mod_lt; auto with zarith.
+intros a Ha.
+pattern a at 1; rewrite <- (Zmod_small a n); auto with zarith.
+repeat rewrite <- Zplus_mod; auto with zarith.
+rewrite <- (Zmod_small 0 n); auto with zarith.
+eq_tac; auto with zarith.
+apply in_mkZp; auto.
+intros a b Ha Hb H; case (prime_mult n H1 a b).
+apply Zmod_divide; auto with zarith.
+intros H2; left.
+case (Zle_lt_or_eq 0 a); auto.
+case (in_mkZp a); auto.
+intros H3; absurd (n <= a).
+apply Zlt_not_le.
+case (in_mkZp a); auto.
+apply Zdivide_le; auto with zarith.
+intros H2; right.
+case (Zle_lt_or_eq 0 b); auto.
+case (in_mkZp b); auto.
+intros H3; absurd (n <= b).
+apply Zlt_not_le.
+case (in_mkZp b); auto.
+apply Zdivide_le; auto with zarith.
+Qed.
+
+End Zp.
+
+(* Definition of the order (0 for q < 1) *)
+
+Definition Zorder: Z -> Z -> Z.
+intros p q; case (Z_le_dec q 1); intros H.
+exact 0.
+refine (e_order Z_eq_dec (p mod q) (ZPGroup q _)); auto with zarith.
+Defined.
+
+Theorem Zorder_pos: forall p n, 0 <= Zorder p n.
+intros p n; unfold Zorder.
+case (Z_le_dec n 1); auto with zarith.
+intros n1.
+apply Zlt_le_weak; apply e_order_pos.
+Qed.
+
+Theorem in_mod_ZPGroup
+ : forall (n : Z) (n_pos : 1 < n) (p : Z),
+ rel_prime p n -> In (p mod n) (s (ZPGroup n n_pos)).
+intros n H p H1.
+apply in_ZPGroup; auto.
+apply rel_prime_mod; auto with zarith.
+apply Z_mod_lt; auto with zarith.
+Qed.
+
+
+Theorem Zpower_mod_is_gpow:
+ forall p q n (Hn: 1 < n), rel_prime p n -> 0 <= q -> p ^ q mod n = gpow (p mod n) (ZPGroup n Hn) q.
+intros p q n H Hp H1; generalize H1; pattern q; apply natlike_ind; simpl; auto.
+intros _; apply Zmod_small; auto with zarith.
+intros n1 Hn1 Rec _; simpl.
+generalize (in_mod_ZPGroup _ H _ Hp); intros Hu.
+unfold Zsucc; rewrite Zpower_exp; try rewrite Zpower_1_r; auto with zarith.
+rewrite gpow_add; auto with zarith.
+rewrite gpow_1; auto; rewrite <- Rec; auto.
+rewrite Zmult_mod; auto.
+Qed.
+
+
+Theorem Zorder_div_power: forall p q n, 1 < n -> rel_prime p n -> p ^ q mod n = 1 -> (Zorder p n | q).
+intros p q n H H1 H2.
+assert (Hq: 0 <= q).
+generalize H2; case q; simpl; auto with zarith.
+intros p1 H3; contradict H3; rewrite Zmod_small; auto with zarith.
+unfold Zorder; case (Z_le_dec n 1).
+intros H3; contradict H; auto with zarith.
+intros H3; apply e_order_divide_gpow; auto.
+apply in_mod_ZPGroup; auto.
+rewrite <- Zpower_mod_is_gpow; auto with zarith.
+Qed.
+
+Theorem Zorder_div: forall p n, prime n -> ~(n | p) -> (Zorder p n | n - 1).
+intros p n H; unfold Zorder.
+case (Z_le_dec n 1); intros H1 H2.
+contradict H1; generalize (prime_ge_2 n H); auto with zarith.
+rewrite <- prime_phi_n_minus_1; auto.
+match goal with |- context[ZPGroup _ ?H2] => rewrite phi_is_order with (n_pos := H2) end.
+apply e_order_divide_g_order; auto.
+apply in_mod_ZPGroup; auto.
+apply rel_prime_sym; apply prime_rel_prime; auto.
+Qed.
+
+
+Theorem Zorder_power_is_1: forall p n, 1 < n -> rel_prime p n -> p ^ (Zorder p n) mod n = 1.
+intros p n H H1; unfold Zorder.
+case (Z_le_dec n 1); intros H2.
+contradict H; auto with zarith.
+let x := match goal with |- context[ZPGroup _ ?X] => X end in rewrite Zpower_mod_is_gpow with (Hn := x); auto with zarith.
+rewrite gpow_e_order_is_e.
+reflexivity.
+apply in_mod_ZPGroup; auto.
+apply Zlt_le_weak; apply e_order_pos.
+Qed.
+
+Theorem Zorder_power_pos: forall p n, 1 < n -> rel_prime p n -> 0 < Zorder p n.
+intros p n H H1; unfold Zorder.
+case (Z_le_dec n 1); intros H2.
+contradict H; auto with zarith.
+apply e_order_pos.
+Qed.
+
+Theorem phi_power_is_1: forall p n, 1 < n -> rel_prime p n -> p ^ (phi n) mod n = 1.
+intros p n H H1.
+assert (V1:= Zorder_power_pos p n H H1).
+assert (H2: (Zorder p n | phi n)).
+unfold Zorder.
+case (Z_le_dec n 1); intros H2.
+contradict H; auto with zarith.
+match goal with |- context[ZPGroup n ?H] =>
+rewrite phi_is_order with (n_pos := H)
+end.
+apply e_order_divide_g_order.
+apply in_mod_ZPGroup; auto.
+case H2; clear H2; intros q H2; rewrite H2.
+rewrite Zmult_comm.
+assert (V2 := (phi_pos _ H)).
+assert (V3: 0 <= q).
+rewrite H2 in V2.
+apply Zlt_le_weak; apply Zmult_lt_0_reg_r with (2 := V2); auto with zarith.
+rewrite Zpower_mult; auto with zarith.
+rewrite Zpower_mod; auto with zarith.
+rewrite Zorder_power_is_1; auto.
+rewrite Zpower_1_l; auto with zarith.
+apply Zmod_small; auto with zarith.
+Qed.