@@ -162,6 +162,7 @@ Compute the_exponent (F2C F9).
162162
163163
164164
165+
165166(** A first attempt to define Fchain correctness *)
166167
167168(* begin snippet BadDefa *)
@@ -1369,9 +1370,9 @@ Function chain_gen (s:signature) {measure signature_measure}
13691370 | (q, 0%N) =>
13701371 Fcompose (chain_gen (gen_F (gamma i)))
13711372 (chain_gen (gen_F (N2pos q)))
1372- | (q,r ) => KFF (chain_gen
1373- (gen_K (N2pos r )
1374- (gamma i - N2pos r )))
1373+ | (q,_r ) => KFF (chain_gen
1374+ (gen_K (N2pos _r )
1375+ (gamma i - N2pos _r )))
13751376 (chain_gen (gen_F (N2pos q)))
13761377 end end
13771378 | gen_K p d =>
@@ -1380,8 +1381,8 @@ Function chain_gen (s:signature) {measure signature_measure}
13801381 match N.pos_div_eucl (p + d) (Npos p) with
13811382 | (q, 0%N) => FFK (chain_gen (gen_F p))
13821383 (chain_gen (gen_F (N2pos q)))
1383- | (q,r ) => KFK (chain_gen (gen_K (N2pos r )
1384- (p - N2pos r )))
1384+ | (q, _r ) => KFK (chain_gen (gen_K (N2pos _r )
1385+ (p - N2pos _r )))
13851386 (chain_gen (gen_F (N2pos q)))
13861387 end
13871388 end .
@@ -1426,23 +1427,23 @@ Proof.
14261427
14271428
14281429 - pattern i at 1;
1429- replace i with (gamma i * (N2pos q) + N2pos r ).
1430+ replace i with (gamma i * (N2pos q) + N2pos _r ).
14301431 + destruct IHc, IHc0;split.
14311432 * apply KFF_correct;auto.
14321433 simpl; simpl in H.
14331434 replace (gamma i) with
1434- (N2pos r + (gamma i - N2pos r )) at 1.
1435+ (N2pos _r + (gamma i - N2pos _r )) at 1.
14351436 apply H.
14361437 rewrite Pplus_minus;auto with chains.
14371438 apply Pos.lt_gt; rewrite N2pos_lt_switch2.
14381439 generalize
14391440 (N.pos_div_eucl_remainder i (N.pos (gamma i) ));
14401441 rewrite e3; simpl;auto with chains.
1441- destruct r ; [ contradiction | auto with chains].
1442+ destruct _r ; [ contradiction | auto with chains].
14421443 * apply KFF_proper;auto with chains.
14431444
14441445 + apply N_pos_div_eucl_rest; auto with chains.
1445- destruct r ;try contradiction; auto with chains.
1446+ destruct _r ;try contradiction; auto with chains.
14461447 apply (div_gamma_pos _ _ _ e3); auto with chains.
14471448 apply pos_gt_3;auto with chains.
14481449 destruct (exact_log2 i); [contradiction | reflexivity].
@@ -1468,15 +1469,15 @@ Proof.
14681469 + apply FFK_proper;auto with chains.
14691470
14701471 - destruct IHc, IHc0; split.
1471- + red; replace (p+d) with (p * N2pos q + N2pos r ).
1472+ + red; replace (p+d) with (p * N2pos q + N2pos _r ).
14721473 * apply KFK_correct;auto with chains.
1473- red in H; replace (N2pos r + (p - N2pos r ))%positive with p in H.
1474+ red in H; replace (N2pos _r + (p - N2pos _r ))%positive with p in H.
14741475 apply H.
14751476 rewrite Pplus_minus; auto.
14761477 generalize (N.pos_div_eucl_remainder (p + d) (N.pos p));
14771478 rewrite e1; cbn; intro H3.
14781479 apply Pos.lt_gt; rewrite N2pos_lt_switch2;auto with chains.
1479- destruct r ; [contradiction | auto with chains].
1480+ destruct _r ; [contradiction | auto with chains].
14801481
14811482 * generalize (N.pos_div_eucl_spec (p + d) (N.pos p));
14821483 rewrite e1; intros H3; clear H H0 H1 H2.
@@ -1486,7 +1487,7 @@ Proof.
14861487 rewrite pos2N_inj_add; apply N.le_add_r.
14871488 }
14881489 {
1489- intros p0 Hp0;subst q; cbn; destruct r ; [ contradiction | ].
1490+ intros p0 Hp0;subst q; cbn; destruct _r ; [ contradiction | ].
14901491 simpl; simpl in H3; injection H3.
14911492 rewrite Pos.mul_comm; auto with chains.
14921493 }
@@ -1508,23 +1509,26 @@ Arguments make_chain gamma {_} _ _ _ .
15081509Compute the_exponent (make_chain dicho 87).
15091510(* end snippet C87Dicho *)
15101511
1511- (** cf Coq workshop 2014 by Jason Grosss *)
15121512
1513- Module Examples.
1513+ Require Import Int63.Uint63.
1514+ Require Import Monoid_instances.
15141515
1515- Import Int31.
1516+ Check cpower (make_chain dicho) 10.
1517+ Module Examples.
15161518Compute cpower (make_chain dicho) 10 12.
15171519Compute cpower (make_chain dicho) 87 12.
1520+ Search (int -> Z).
1521+ Search (positive -> int).
1522+
15181523
1519- Definition fast_int31_power (x :positive)(n:N) : Z :=
1520- Int31.phi (cpower (make_chain dicho) n (snd (positive_to_int31 x) )).
1524+ Definition fast_int63_power (x :positive)(n:N) : Z :=
1525+ to_Z (cpower (make_chain dicho) n (of_pos x )).
15211526
15221527Definition slow_int31_power (x :positive)(n:N) : Z :=
1523- Int31.phi (power (snd (positive_to_int31 x) ) (N.to_nat n) ).
1528+ to_Z (power (of_pos x ) (N.to_nat n) ).
15241529
15251530Definition binary_int31_power (x :positive)(n:N) : Z :=
1526- Int31.phi (N_bpow (snd (positive_to_int31 x)) n ).
1527-
1531+ to_Z (N_bpow (of_pos x) n).
15281532
15291533(** long computations ... *)
15301534
@@ -1537,7 +1541,7 @@ Arguments big_chain _%type_scope.
15371541Remark RM : (1 < 56789)%N.
15381542Proof . reflexivity. Qed .
15391543
1540- Definition M := Nmod_Monoid _ RM.
1544+ Definition M := Nmod_Monoid 56789%N RM.
15411545
15421546Definition exp56789 x := chain_apply big_chain (M:=M) x.
15431547
@@ -1550,12 +1554,6 @@ Eval cbv iota match delta [big_chain chain_apply computation_eval ] zeta beta
15501554Definition C87' := ltac:( compute_chain C87 ).
15511555
15521556
1553- Time Compute Int31.phi
1554- (chain_apply big_chain (snd (positive_to_int31 67777))) .
1555-
1556-
1557- Compute Int31.phi (chain_apply big_chain (snd (positive_to_int31 67777))) .
1558-
15591557Compute chain_length big_chain.
15601558
15611559Goal parametric (make_chain dicho 45319).
0 commit comments