aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar thery <thery@85f007b7-540e-0410-9357-904b9bb8a0f7>2007-10-02 11:42:42 +0000
committerGravatar thery <thery@85f007b7-540e-0410-9357-904b9bb8a0f7>2007-10-02 11:42:42 +0000
commit95d72c0eb4a7adc28a277272d1a88b0db453a9d4 (patch)
tree4ec77c3b609516911ca9afe8186c8d91fd692d5c
parent5a419bafdb21fd61c4c5c0784c1a0156d5ec7005 (diff)
Now NMake is proved
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@10163 85f007b7-540e-0410-9357-904b9bb8a0f7
-rw-r--r--theories/Ints/num/GenBase.v96
-rw-r--r--theories/Ints/num/GenLift.v62
-rw-r--r--theories/Ints/num/NMake.v5639
-rw-r--r--theories/Ints/num/Zn2Z.v25
-rw-r--r--theories/Ints/num/ZnZ.v2
-rw-r--r--theories/Ints/num/genN.ml4471
6 files changed, 5334 insertions, 4961 deletions
diff --git a/theories/Ints/num/GenBase.v b/theories/Ints/num/GenBase.v
index 04749f79c..6b6376b69 100644
--- a/theories/Ints/num/GenBase.v
+++ b/theories/Ints/num/GenBase.v
@@ -132,6 +132,21 @@ Section GenBase.
| Gt => Gt
end
end.
+
+
+ (* Return the low part of the composed word*)
+ Fixpoint get_low (n : nat) {struct n}:
+ word w n -> w :=
+ match n return (word w n -> w) with
+ | 0%nat => fun x => x
+ | S n1 =>
+ fun x =>
+ match x with
+ | W0 => w_0
+ | WW _ x1 => get_low n1 x1
+ end
+ end.
+
Section GenProof.
Notation wB := (base w_digits).
@@ -288,24 +303,72 @@ Section GenBase.
ring.
Qed.
+ Lemma gen_wB_pos:
+ forall n, 0 <= gen_wB n.
+ Proof.
+ intros n; unfold gen_wB, base; auto with zarith.
+ Qed.
+
+ Lemma gen_wB_more_digits:
+ forall n, wB <= gen_wB n.
+ Proof.
+ clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1.
+ intros n; elim n; clear n; auto.
+ unfold gen_wB, gen_digits; auto with zarith.
+ intros n H1; rewrite <- gen_wB_wwB.
+ apply Zle_trans with (wB * 1).
+ rewrite Zmult_1_r; apply Zle_refl.
+ apply Zmult_le_compat; auto with zarith.
+ apply Zle_trans with wB; auto with zarith.
+ unfold base.
+ rewrite <- (ZAux.Zpower_exp_0 2).
+ apply Zpower_le_monotone2; auto with zarith.
+ unfold base; auto with zarith.
+ Qed.
+
Lemma spec_gen_to_Z :
- forall n (x:word w n), 0 <= gen_to_Z n x < gen_wB n.
+ forall n (x:word w n), 0 <= [!n | x!] < gen_wB n.
Proof.
- clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1.
- induction n;intros. exact (spec_to_Z x).
- unfold gen_to_Z;fold gen_to_Z.
- destruct x;unfold zn2z_to_Z.
- unfold gen_wB,base;split;auto with zarith.
- assert (U0:= IHn w0);assert (U1:= IHn w1).
- split;auto with zarith.
- apply Zlt_le_trans with ((gen_wB n - 1) * gen_wB n + gen_wB n).
- assert (gen_to_Z n w0*gen_wB n <= (gen_wB n - 1)*gen_wB n).
- apply Zmult_le_compat_r;auto with zarith.
- auto with zarith.
- rewrite <- gen_wB_wwB.
- replace ((gen_wB n - 1) * gen_wB n + gen_wB n) with (gen_wB n * gen_wB n);
- [auto with zarith | ring].
- Qed.
+ clear spec_w_0 spec_w_1 spec_w_Bm1 w_0 w_1 w_Bm1.
+ induction n;intros. exact (spec_to_Z x).
+ unfold gen_to_Z;fold gen_to_Z.
+ destruct x;unfold zn2z_to_Z.
+ unfold gen_wB,base;split;auto with zarith.
+ assert (U0:= IHn w0);assert (U1:= IHn w1).
+ split;auto with zarith.
+ apply Zlt_le_trans with ((gen_wB n - 1) * gen_wB n + gen_wB n).
+ assert (gen_to_Z n w0*gen_wB n <= (gen_wB n - 1)*gen_wB n).
+ apply Zmult_le_compat_r;auto with zarith.
+ auto with zarith.
+ rewrite <- gen_wB_wwB.
+ replace ((gen_wB n - 1) * gen_wB n + gen_wB n) with (gen_wB n * gen_wB n);
+ [auto with zarith | ring].
+ Qed.
+
+ Lemma spec_get_low:
+ forall n x,
+ [!n | x!] < wB -> [|get_low n x|] = [!n | x!].
+ Proof.
+ clear spec_w_1 spec_w_Bm1.
+ intros n; elim n; auto; clear n.
+ intros n Hrec x; case x; clear x; auto.
+ intros xx yy H1; simpl in H1.
+ assert (F1: [!n | xx!] = 0).
+ case (Zle_lt_or_eq 0 ([!n | xx!])); auto.
+ case (spec_gen_to_Z n xx); auto.
+ intros F2.
+ assert (F3 := gen_wB_more_digits n).
+ assert (F4: 0 <= [!n | yy!]).
+ case (spec_gen_to_Z n yy); auto.
+ assert (F5: 1 * wB <= [!n | xx!] * gen_wB n);
+ auto with zarith.
+ apply Zmult_le_compat; auto with zarith.
+ unfold base; auto with zarith.
+ simpl get_low; simpl gen_to_Z.
+ generalize H1; clear H1.
+ rewrite F1; rewrite Zmult_0_l; rewrite Zplus_0_l.
+ intros H1; apply Hrec; auto.
+ Qed.
Lemma spec_gen_WW : forall n (h l : word w n),
[!S n|gen_WW n h l!] = [!n|h!] * gen_wB n + [!n|l!].
@@ -374,6 +437,7 @@ Section GenBase.
apply wB_lex_inv;trivial.
apply Zlt_gt;apply wB_lex_inv;apply Zgt_lt;trivial.
Qed.
+
End GenProof.
diff --git a/theories/Ints/num/GenLift.v b/theories/Ints/num/GenLift.v
index c1283aefb..11455b804 100644
--- a/theories/Ints/num/GenLift.v
+++ b/theories/Ints/num/GenLift.v
@@ -117,8 +117,10 @@ Section GenLift.
| Gt => [[x]] > [[y]]
end.
Variable spec_ww_digits : ww_Digits = xO w_digits.
+ Variable spec_w_head00 : forall x, [|x|] = 0 -> [|w_head0 x|] = Zpos w_digits.
Variable spec_w_head0 : forall x, 0 < [|x|] ->
wB/ 2 <= 2 ^ ([|w_head0 x|]) * [|x|] < wB.
+ Variable spec_w_tail00 : forall x, [|x|] = 0 -> [|w_tail0 x|] = Zpos w_digits.
Variable spec_w_tail0 : forall x, 0 < [|x|] ->
exists y, 0 <= y /\ [|x|] = (2* y + 1) * (2 ^ [|w_tail0 x|]).
Variable spec_w_add_mul_div : forall x y p,
@@ -134,13 +136,42 @@ Section GenLift.
Variable spec_zdigits : [| w_zdigits |] = Zpos w_digits.
Variable spec_low: forall x, [| low x|] = [[x]] mod wB.
+ Variable spec_ww_zdigits : [[ww_zdigits]] = Zpos ww_Digits.
Hint Resolve div_le_0 div_lt w_to_Z_wwB: lift.
Ltac zarith := auto with zarith lift.
+
+ Lemma spec_ww_head00 : forall x, [[x]] = 0 -> [[ww_head0 x]] = Zpos ww_Digits.
+ Proof.
+ intros x; case x; unfold ww_head0.
+ intros HH; rewrite spec_ww_zdigits; auto.
+ intros xh xl; simpl; intros Hx.
+ case (spec_to_Z xh); intros Hx1 Hx2.
+ case (spec_to_Z xl); intros Hy1 Hy2.
+ assert (F1: [|xh|] = 0).
+ case (Zle_lt_or_eq _ _ Hy1); auto; intros Hy3.
+ absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith.
+ apply Zlt_le_trans with (1 := Hy3); auto with zarith.
+ pattern [|xl|] at 1; rewrite <- (Zplus_0_l [|xl|]).
+ apply Zplus_le_compat_r; auto with zarith.
+ case (Zle_lt_or_eq _ _ Hx1); auto; intros Hx3.
+ absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith.
+ rewrite <- Hy3; rewrite Zplus_0_r; auto with zarith.
+ apply Zmult_lt_0_compat; auto with zarith.
+ generalize (spec_compare w_0 xh); case w_compare.
+ intros H; simpl.
+ rewrite spec_w_add; rewrite spec_w_head00.
+ rewrite spec_zdigits; rewrite spec_ww_digits.
+ rewrite Zpos_xO; auto with zarith.
+ rewrite F1 in Hx; auto with zarith.
+ rewrite spec_w_0; auto with zarith.
+ rewrite spec_w_0; auto with zarith.
+ Qed.
Lemma spec_ww_head0 : forall x, 0 < [[x]] ->
wwB/ 2 <= 2 ^ [[ww_head0 x]] * [[x]] < wwB.
Proof.
+ clear spec_ww_zdigits.
rewrite wwB_div_2;rewrite Zmult_comm;rewrite wwB_wBwB.
assert (U:= lt_0_wB w_digits); destruct x as [ |xh xl];simpl ww_to_Z;intros H.
unfold Zlt in H;discriminate H.
@@ -182,9 +213,38 @@ Section GenLift.
assert (H1 := spec_to_Z xh);zarith.
Qed.
+ Lemma spec_ww_tail00 : forall x, [[x]] = 0 -> [[ww_tail0 x]] = Zpos ww_Digits.
+ Proof.
+ intros x; case x; unfold ww_tail0.
+ intros HH; rewrite spec_ww_zdigits; auto.
+ intros xh xl; simpl; intros Hx.
+ case (spec_to_Z xh); intros Hx1 Hx2.
+ case (spec_to_Z xl); intros Hy1 Hy2.
+ assert (F1: [|xh|] = 0).
+ case (Zle_lt_or_eq _ _ Hy1); auto; intros Hy3.
+ absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith.
+ apply Zlt_le_trans with (1 := Hy3); auto with zarith.
+ pattern [|xl|] at 1; rewrite <- (Zplus_0_l [|xl|]).
+ apply Zplus_le_compat_r; auto with zarith.
+ case (Zle_lt_or_eq _ _ Hx1); auto; intros Hx3.
+ absurd (0 < [|xh|] * wB + [|xl|]); auto with zarith.
+ rewrite <- Hy3; rewrite Zplus_0_r; auto with zarith.
+ apply Zmult_lt_0_compat; auto with zarith.
+ assert (F2: [|xl|] = 0).
+ rewrite F1 in Hx; auto with zarith.
+ generalize (spec_compare w_0 xl); case w_compare.
+ intros H; simpl.
+ rewrite spec_w_add; rewrite spec_w_tail00; auto.
+ rewrite spec_zdigits; rewrite spec_ww_digits.
+ rewrite Zpos_xO; auto with zarith.
+ rewrite spec_w_0; auto with zarith.
+ rewrite spec_w_0; auto with zarith.
+ Qed.
+
Lemma spec_ww_tail0 : forall x, 0 < [[x]] ->
exists y, 0 <= y /\ [[x]] = (2 * y + 1) * 2 ^ [[ww_tail0 x]].
Proof.
+ clear spec_ww_zdigits.
destruct x as [ |xh xl];simpl ww_to_Z;intros H.
unfold Zlt in H;discriminate H.
assert (H0 := spec_compare w_0 xl);rewrite spec_w_0 in H0.
@@ -251,6 +311,7 @@ Section GenLift.
([[WW xh xl]] * (2^[[p]]) +
[[WW yh yl]] / (2^(Zpos (xO w_digits) - [[p]]))) mod wwB.
Proof.
+ clear spec_ww_zdigits.
intros xh xl yh yl p zdigits;assert (HwwB := wwB_pos w_digits).
case (spec_to_w_Z p); intros Hv1 Hv2.
replace (Zpos (xO w_digits)) with (Zpos w_digits + Zpos w_digits).
@@ -374,6 +435,7 @@ Section GenLift.
([[x]] * (2^[[p]]) +
[[y]] / (2^(Zpos (xO w_digits) - [[p]]))) mod wwB.
Proof.
+ clear spec_ww_zdigits.
intros x y p H.
destruct x as [ |xh xl];
[assert (H1 := @spec_ww_add_mul_div_aux w_0 w_0)
diff --git a/theories/Ints/num/NMake.v b/theories/Ints/num/NMake.v
index 26174bdb1..03402c53f 100644
--- a/theories/Ints/num/NMake.v
+++ b/theories/Ints/num/NMake.v
@@ -1,3 +1,4 @@
+Require Import ZAux.
Require Import ZArith.
Require Import Basic_type.
Require Import ZnZ.
@@ -7,6 +8,13 @@ Require Import GenMul.
Require Import GenDivn1.
Require Import Wf_nat.
+ (***************************************************************)
+ (* *)
+ (* File automatically generated DO NOT EDIT *)
+ (* Constructors: 13 Generated Proofs: false *)
+ (* *)
+ (***************************************************************)
+
Module Type W0Type.
Parameter w : Set.
Parameter w_op : znz_op w.
@@ -109,6 +117,1508 @@ Module Make (W0:W0Type).
Definition zero := N0 w_0.
Definition one := N0 one0.
+ Definition to_Z x :=
+ match x with
+ | N0 wx => w0_op.(znz_to_Z) wx
+ | N1 wx => w1_op.(znz_to_Z) wx
+ | N2 wx => w2_op.(znz_to_Z) wx
+ | N3 wx => w3_op.(znz_to_Z) wx
+ | N4 wx => w4_op.(znz_to_Z) wx
+ | N5 wx => w5_op.(znz_to_Z) wx
+ | N6 wx => w6_op.(znz_to_Z) wx
+ | N7 wx => w7_op.(znz_to_Z) wx
+ | N8 wx => w8_op.(znz_to_Z) wx
+ | N9 wx => w9_op.(znz_to_Z) wx
+ | N10 wx => w10_op.(znz_to_Z) wx
+ | N11 wx => w11_op.(znz_to_Z) wx
+ | N12 wx => w12_op.(znz_to_Z) wx
+ | N13 wx => w13_op.(znz_to_Z) wx
+ | Nn n wx => (make_op n).(znz_to_Z) wx
+ end.
+
+ Open Scope Z_scope.
+ Notation "[ x ]" := (to_Z x).
+
+ (* Eval and extend functions for each level *)
+ Let extend0 := GenBase.extend (WW w_0).
+ Let extend1 := GenBase.extend (WW (W0: w1)).
+ Let extend2 := GenBase.extend (WW (W0: w2)).
+ Let extend3 := GenBase.extend (WW (W0: w3)).
+ Let extend4 := GenBase.extend (WW (W0: w4)).
+ Let extend5 := GenBase.extend (WW (W0: w5)).
+ Let extend6 := GenBase.extend (WW (W0: w6)).
+ Let extend7 := GenBase.extend (WW (W0: w7)).
+ Let extend8 := GenBase.extend (WW (W0: w8)).
+ Let extend9 := GenBase.extend (WW (W0: w9)).
+ Let extend10 := GenBase.extend (WW (W0: w10)).
+ Let extend11 := GenBase.extend (WW (W0: w11)).
+ Let extend12 := GenBase.extend (WW (W0: w12)).
+ Let extend13 := GenBase.extend (WW (W0: w13)).
+
+ Definition w0_eq0 := w0_op.(znz_eq0).
+ Let spec_w0_eq0: forall x, if w0_eq0 x then [N0 x] = 0 else True.
+ Admitted.
+
+ Definition w1_eq0 := w1_op.(znz_eq0).
+ Let spec_w1_eq0: forall x, if w1_eq0 x then [N1 x] = 0 else True.
+ Admitted.
+
+ Definition w2_eq0 := w2_op.(znz_eq0).
+ Let spec_w2_eq0: forall x, if w2_eq0 x then [N2 x] = 0 else True.
+ Admitted.
+
+ Definition w3_eq0 := w3_op.(znz_eq0).
+ Let spec_w3_eq0: forall x, if w3_eq0 x then [N3 x] = 0 else True.
+ Admitted.
+
+ Definition w4_eq0 := w4_op.(znz_eq0).
+ Let spec_w4_eq0: forall x, if w4_eq0 x then [N4 x] = 0 else True.
+ Admitted.
+
+ Definition w5_eq0 := w5_op.(znz_eq0).
+ Let spec_w5_eq0: forall x, if w5_eq0 x then [N5 x] = 0 else True.
+ Admitted.
+
+ Definition w6_eq0 := w6_op.(znz_eq0).
+ Let spec_w6_eq0: forall x, if w6_eq0 x then [N6 x] = 0 else True.
+ Admitted.
+
+ Definition w7_eq0 := w7_op.(znz_eq0).
+ Let spec_w7_eq0: forall x, if w7_eq0 x then [N7 x] = 0 else True.
+ Admitted.
+
+ Definition w8_eq0 := w8_op.(znz_eq0).
+ Let spec_w8_eq0: forall x, if w8_eq0 x then [N8 x] = 0 else True.
+ Admitted.
+
+ Definition w9_eq0 := w9_op.(znz_eq0).
+ Let spec_w9_eq0: forall x, if w9_eq0 x then [N9 x] = 0 else True.
+ Admitted.
+
+ Definition w10_eq0 := w10_op.(znz_eq0).
+ Let spec_w10_eq0: forall x, if w10_eq0 x then [N10 x] = 0 else True.
+ Admitted.
+
+ Definition w11_eq0 := w11_op.(znz_eq0).
+ Let spec_w11_eq0: forall x, if w11_eq0 x then [N11 x] = 0 else True.
+ Admitted.
+
+ Definition w12_eq0 := w12_op.(znz_eq0).
+ Let spec_w12_eq0: forall x, if w12_eq0 x then [N12 x] = 0 else True.
+ Admitted.
+
+ Definition w13_eq0 := w13_op.(znz_eq0).
+ Let spec_w13_eq0: forall x, if w13_eq0 x then [N13 x] = 0 else True.
+ Admitted.
+
+
+ Theorem to_Z_pos: forall x, 0 <= [x].
+ Admitted.
+
+ Section LevelAndIter.
+
+ Variable res: Set.
+ Variable xxx: res.
+ Variable P: Z -> Z -> res -> Prop.
+ (* Abstraction function for each level *)
+ Variable f0: w0 -> w0 -> res.
+ Variable f0n: forall n, w0 -> word w0 (S n) -> res.
+ Variable fn0: forall n, word w0 (S n) -> w0 -> res.
+
+ Variable f1: w1 -> w1 -> res.
+ Variable f1n: forall n, w1 -> word w1 (S n) -> res.
+ Variable fn1: forall n, word w1 (S n) -> w1 -> res.
+
+ Variable f2: w2 -> w2 -> res.
+ Variable f2n: forall n, w2 -> word w2 (S n) -> res.
+ Variable fn2: forall n, word w2 (S n) -> w2 -> res.
+
+ Variable f3: w3 -> w3 -> res.
+ Variable f3n: forall n, w3 -> word w3 (S n) -> res.
+ Variable fn3: forall n, word w3 (S n) -> w3 -> res.
+
+ Variable f4: w4 -> w4 -> res.
+ Variable f4n: forall n, w4 -> word w4 (S n) -> res.
+ Variable fn4: forall n, word w4 (S n) -> w4 -> res.
+
+ Variable f5: w5 -> w5 -> res.
+ Variable f5n: forall n, w5 -> word w5 (S n) -> res.
+ Variable fn5: forall n, word w5 (S n) -> w5 -> res.
+
+ Variable f6: w6 -> w6 -> res.
+ Variable f6n: forall n, w6 -> word w6 (S n) -> res.
+ Variable fn6: forall n, word w6 (S n) -> w6 -> res.
+
+ Variable f7: w7 -> w7 -> res.
+ Variable f7n: forall n, w7 -> word w7 (S n) -> res.
+ Variable fn7: forall n, word w7 (S n) -> w7 -> res.
+
+ Variable f8: w8 -> w8 -> res.
+ Variable f8n: forall n, w8 -> word w8 (S n) -> res.
+ Variable fn8: forall n, word w8 (S n) -> w8 -> res.
+
+ Variable f9: w9 -> w9 -> res.
+ Variable f9n: forall n, w9 -> word w9 (S n) -> res.
+ Variable fn9: forall n, word w9 (S n) -> w9 -> res.
+
+ Variable f10: w10 -> w10 -> res.
+ Variable f10n: forall n, w10 -> word w10 (S n) -> res.
+ Variable fn10: forall n, word w10 (S n) -> w10 -> res.
+
+ Variable f11: w11 -> w11 -> res.
+ Variable f11n: forall n, w11 -> word w11 (S n) -> res.
+ Variable fn11: forall n, word w11 (S n) -> w11 -> res.
+
+ Variable f12: w12 -> w12 -> res.
+ Variable f12n: forall n, w12 -> word w12 (S n) -> res.
+ Variable fn12: forall n, word w12 (S n) -> w12 -> res.
+
+ Variable f13: w13 -> w13 -> res.
+ Variable f13n: forall n, w13 -> word w13 (S n) -> res.
+ Variable fn13: forall n, word w13 (S n) -> w13 -> res.
+
+ Variable fnn: forall n, word w13 (S n) -> word w13 (S n) -> res.
+ Variable fnm: forall n m, word w13 (S n) -> word w13 (S m) -> res.
+
+ (* Special zero functions *)
+ Variable f0t: t_ -> res.
+ Variable ft0: t_ -> res.
+
+ (* We level the two arguments before applying *)
+ (* the functions at each leval *)
+ Definition same_level (x y: t_): res :=
+ Eval lazy zeta beta iota delta [extend0 extend1 extend2 extend3 extend4 extend5 extend6 extend7 extend8 extend9 extend10 extend11 extend12 extend13
+ GenBase.extend GenBase.extend_aux
+ ] in
+ match x, y with
+ | N0 wx, N0 wy => f0 wx wy
+ | N0 wx, N1 wy => f1 (extend0 0 wx) wy
+ | N0 wx, N2 wy => f2 (extend0 1 wx) wy
+ | N0 wx, N3 wy => f3 (extend0 2 wx) wy
+ | N0 wx, N4 wy => f4 (extend0 3 wx) wy
+ | N0 wx, N5 wy => f5 (extend0 4 wx) wy
+ | N0 wx, N6 wy => f6 (extend0 5 wx) wy
+ | N0 wx, N7 wy => f7 (extend0 6 wx) wy
+ | N0 wx, N8 wy => f8 (extend0 7 wx) wy
+ | N0 wx, N9 wy => f9 (extend0 8 wx) wy
+ | N0 wx, N10 wy => f10 (extend0 9 wx) wy
+ | N0 wx, N11 wy => f11 (extend0 10 wx) wy
+ | N0 wx, N12 wy => f12 (extend0 11 wx) wy
+ | N0 wx, N13 wy => f13 (extend0 12 wx) wy
+ | N0 wx, Nn m wy => fnn m (extend13 m (extend0 12 wx)) wy
+ | N1 wx, N0 wy => f1 wx (extend0 0 wy)
+ | N1 wx, N1 wy => f1 wx wy
+ | N1 wx, N2 wy => f2 (extend1 0 wx) wy
+ | N1 wx, N3 wy => f3 (extend1 1 wx) wy
+ | N1 wx, N4 wy => f4 (extend1 2 wx) wy
+ | N1 wx, N5 wy => f5 (extend1 3 wx) wy
+ | N1 wx, N6 wy => f6 (extend1 4 wx) wy
+ | N1 wx, N7 wy => f7 (extend1 5 wx) wy
+ | N1 wx, N8 wy => f8 (extend1 6 wx) wy
+ | N1 wx, N9 wy => f9 (extend1 7 wx) wy
+ | N1 wx, N10 wy => f10 (extend1 8 wx) wy
+ | N1 wx, N11 wy => f11 (extend1 9 wx) wy
+ | N1 wx, N12 wy => f12 (extend1 10 wx) wy
+ | N1 wx, N13 wy => f13 (extend1 11 wx) wy
+ | N1 wx, Nn m wy => fnn m (extend13 m (extend1 11 wx)) wy
+ | N2 wx, N0 wy => f2 wx (extend0 1 wy)
+ | N2 wx, N1 wy => f2 wx (extend1 0 wy)
+ | N2 wx, N2 wy => f2 wx wy
+ | N2 wx, N3 wy => f3 (extend2 0 wx) wy
+ | N2 wx, N4 wy => f4 (extend2 1 wx) wy
+ | N2 wx, N5 wy => f5 (extend2 2 wx) wy
+ | N2 wx, N6 wy => f6 (extend2 3 wx) wy
+ | N2 wx, N7 wy => f7 (extend2 4 wx) wy
+ | N2 wx, N8 wy => f8 (extend2 5 wx) wy
+ | N2 wx, N9 wy => f9 (extend2 6 wx) wy
+ | N2 wx, N10 wy => f10 (extend2 7 wx) wy
+ | N2 wx, N11 wy => f11 (extend2 8 wx) wy
+ | N2 wx, N12 wy => f12 (extend2 9 wx) wy
+ | N2 wx, N13 wy => f13 (extend2 10 wx) wy
+ | N2 wx, Nn m wy => fnn m (extend13 m (extend2 10 wx)) wy
+ | N3 wx, N0 wy => f3 wx (extend0 2 wy)
+ | N3 wx, N1 wy => f3 wx (extend1 1 wy)
+ | N3 wx, N2 wy => f3 wx (extend2 0 wy)
+ | N3 wx, N3 wy => f3 wx wy
+ | N3 wx, N4 wy => f4 (extend3 0 wx) wy
+ | N3 wx, N5 wy => f5 (extend3 1 wx) wy
+ | N3 wx, N6 wy => f6 (extend3 2 wx) wy
+ | N3 wx, N7 wy => f7 (extend3 3 wx) wy
+ | N3 wx, N8 wy => f8 (extend3 4 wx) wy
+ | N3 wx, N9 wy => f9 (extend3 5 wx) wy
+ | N3 wx, N10 wy => f10 (extend3 6 wx) wy
+ | N3 wx, N11 wy => f11 (extend3 7 wx) wy
+ | N3 wx, N12 wy => f12 (extend3 8 wx) wy
+ | N3 wx, N13 wy => f13 (extend3 9 wx) wy
+ | N3 wx, Nn m wy => fnn m (extend13 m (extend3 9 wx)) wy
+ | N4 wx, N0 wy => f4 wx (extend0 3 wy)
+ | N4 wx, N1 wy => f4 wx (extend1 2 wy)
+ | N4 wx, N2 wy => f4 wx (extend2 1 wy)
+ | N4 wx, N3 wy => f4 wx (extend3 0 wy)
+ | N4 wx, N4 wy => f4 wx wy
+ | N4 wx, N5 wy => f5 (extend4 0 wx) wy
+ | N4 wx, N6 wy => f6 (extend4 1 wx) wy
+ | N4 wx, N7 wy => f7 (extend4 2 wx) wy
+ | N4 wx, N8 wy => f8 (extend4 3 wx) wy
+ | N4 wx, N9 wy => f9 (extend4 4 wx) wy
+ | N4 wx, N10 wy => f10 (extend4 5 wx) wy
+ | N4 wx, N11 wy => f11 (extend4 6 wx) wy
+ | N4 wx, N12 wy => f12 (extend4 7 wx) wy
+ | N4 wx, N13 wy => f13 (extend4 8 wx) wy
+ | N4 wx, Nn m wy => fnn m (extend13 m (extend4 8 wx)) wy
+ | N5 wx, N0 wy => f5 wx (extend0 4 wy)
+ | N5 wx, N1 wy => f5 wx (extend1 3 wy)
+ | N5 wx, N2 wy => f5 wx (extend2 2 wy)
+ | N5 wx, N3 wy => f5 wx (extend3 1 wy)
+ | N5 wx, N4 wy => f5 wx (extend4 0 wy)
+ | N5 wx, N5 wy => f5 wx wy
+ | N5 wx, N6 wy => f6 (extend5 0 wx) wy
+ | N5 wx, N7 wy => f7 (extend5 1 wx) wy
+ | N5 wx, N8 wy => f8 (extend5 2 wx) wy
+ | N5 wx, N9 wy => f9 (extend5 3 wx) wy
+ | N5 wx, N10 wy => f10 (extend5 4 wx) wy
+ | N5 wx, N11 wy => f11 (extend5 5 wx) wy
+ | N5 wx, N12 wy => f12 (extend5 6 wx) wy
+ | N5 wx, N13 wy => f13 (extend5 7 wx) wy
+ | N5 wx, Nn m wy => fnn m (extend13 m (extend5 7 wx)) wy
+ | N6 wx, N0 wy => f6 wx (extend0 5 wy)
+ | N6 wx, N1 wy => f6 wx (extend1 4 wy)
+ | N6 wx, N2 wy => f6 wx (extend2 3 wy)
+ | N6 wx, N3 wy => f6 wx (extend3 2 wy)
+ | N6 wx, N4 wy => f6 wx (extend4 1 wy)
+ | N6 wx, N5 wy => f6 wx (extend5 0 wy)
+ | N6 wx, N6 wy => f6 wx wy
+ | N6 wx, N7 wy => f7 (extend6 0 wx) wy
+ | N6 wx, N8 wy => f8 (extend6 1 wx) wy
+ | N6 wx, N9 wy => f9 (extend6 2 wx) wy
+ | N6 wx, N10 wy => f10 (extend6 3 wx) wy
+ | N6 wx, N11 wy => f11 (extend6 4 wx) wy
+ | N6 wx, N12 wy => f12 (extend6 5 wx) wy
+ | N6 wx, N13 wy => f13 (extend6 6 wx) wy
+ | N6 wx, Nn m wy => fnn m (extend13 m (extend6 6 wx)) wy
+ | N7 wx, N0 wy => f7 wx (extend0 6 wy)
+ | N7 wx, N1 wy => f7 wx (extend1 5 wy)
+ | N7 wx, N2 wy => f7 wx (extend2 4 wy)
+ | N7 wx, N3 wy => f7 wx (extend3 3 wy)
+ | N7 wx, N4 wy => f7 wx (extend4 2 wy)
+ | N7 wx, N5 wy => f7 wx (extend5 1 wy)
+ | N7 wx, N6 wy => f7 wx (extend6 0 wy)
+ | N7 wx, N7 wy => f7 wx wy
+ | N7 wx, N8 wy => f8 (extend7 0 wx) wy
+ | N7 wx, N9 wy => f9 (extend7 1 wx) wy
+ | N7 wx, N10 wy => f10 (extend7 2 wx) wy
+ | N7 wx, N11 wy => f11 (extend7 3 wx) wy
+ | N7 wx, N12 wy => f12 (extend7 4 wx) wy
+ | N7 wx, N13 wy => f13 (extend7 5 wx) wy
+ | N7 wx, Nn m wy => fnn m (extend13 m (extend7 5 wx)) wy
+ | N8 wx, N0 wy => f8 wx (extend0 7 wy)
+ | N8 wx, N1 wy => f8 wx (extend1 6 wy)
+ | N8 wx, N2 wy => f8 wx (extend2 5 wy)
+ | N8 wx, N3 wy => f8 wx (extend3 4 wy)
+ | N8 wx, N4 wy => f8 wx (extend4 3 wy)
+ | N8 wx, N5 wy => f8 wx (extend5 2 wy)
+ | N8 wx, N6 wy => f8 wx (extend6 1 wy)
+ | N8 wx, N7 wy => f8 wx (extend7 0 wy)
+ | N8 wx, N8 wy => f8 wx wy
+ | N8 wx, N9 wy => f9 (extend8 0 wx) wy
+ | N8 wx, N10 wy => f10 (extend8 1 wx) wy
+ | N8 wx, N11 wy => f11 (extend8 2 wx) wy
+ | N8 wx, N12 wy => f12 (extend8 3 wx) wy
+ | N8 wx, N13 wy => f13 (extend8 4 wx) wy
+ | N8 wx, Nn m wy => fnn m (extend13 m (extend8 4 wx)) wy
+ | N9 wx, N0 wy => f9 wx (extend0 8 wy)
+ | N9 wx, N1 wy => f9 wx (extend1 7 wy)
+ | N9 wx, N2 wy => f9 wx (extend2 6 wy)
+ | N9 wx, N3 wy => f9 wx (extend3 5 wy)
+ | N9 wx, N4 wy => f9 wx (extend4 4 wy)
+ | N9 wx, N5 wy => f9 wx (extend5 3 wy)
+ | N9 wx, N6 wy => f9 wx (extend6 2 wy)
+ | N9 wx, N7 wy => f9 wx (extend7 1 wy)
+ | N9 wx, N8 wy => f9 wx (extend8 0 wy)
+ | N9 wx, N9 wy => f9 wx wy
+ | N9 wx, N10 wy => f10 (extend9 0 wx) wy
+ | N9 wx, N11 wy => f11 (extend9 1 wx) wy
+ | N9 wx, N12 wy => f12 (extend9 2 wx) wy
+ | N9 wx, N13 wy => f13 (extend9 3 wx) wy
+ | N9 wx, Nn m wy => fnn m (extend13 m (extend9 3 wx)) wy
+ | N10 wx, N0 wy => f10 wx (extend0 9 wy)
+ | N10 wx, N1 wy => f10 wx (extend1 8 wy)
+ | N10 wx, N2 wy => f10 wx (extend2 7 wy)
+ | N10 wx, N3 wy => f10 wx (extend3 6 wy)
+ | N10 wx, N4 wy => f10 wx (extend4 5 wy)
+ | N10 wx, N5 wy => f10 wx (extend5 4 wy)
+ | N10 wx, N6 wy => f10 wx (extend6 3 wy)
+ | N10 wx, N7 wy => f10 wx (extend7 2 wy)
+ | N10 wx, N8 wy => f10 wx (extend8 1 wy)
+ | N10 wx, N9 wy => f10 wx (extend9 0 wy)
+ | N10 wx, N10 wy => f10 wx wy
+ | N10 wx, N11 wy => f11 (extend10 0 wx) wy
+ | N10 wx, N12 wy => f12 (extend10 1 wx) wy
+ | N10 wx, N13 wy => f13 (extend10 2 wx) wy
+ | N10 wx, Nn m wy => fnn m (extend13 m (extend10 2 wx)) wy
+ | N11 wx, N0 wy => f11 wx (extend0 10 wy)
+ | N11 wx, N1 wy => f11 wx (extend1 9 wy)
+ | N11 wx, N2 wy => f11 wx (extend2 8 wy)
+ | N11 wx, N3 wy => f11 wx (extend3 7 wy)
+ | N11 wx, N4 wy => f11 wx (extend4 6 wy)
+ | N11 wx, N5 wy => f11 wx (extend5 5 wy)
+ | N11 wx, N6 wy => f11 wx (extend6 4 wy)
+ | N11 wx, N7 wy => f11 wx (extend7 3 wy)
+ | N11 wx, N8 wy => f11 wx (extend8 2 wy)
+ | N11 wx, N9 wy => f11 wx (extend9 1 wy)
+ | N11 wx, N10 wy => f11 wx (extend10 0 wy)
+ | N11 wx, N11 wy => f11 wx wy
+ | N11 wx, N12 wy => f12 (extend11 0 wx) wy
+ | N11 wx, N13 wy => f13 (extend11 1 wx) wy
+ | N11 wx, Nn m wy => fnn m (extend13 m (extend11 1 wx)) wy
+ | N12 wx, N0 wy => f12 wx (extend0 11 wy)
+ | N12 wx, N1 wy => f12 wx (extend1 10 wy)
+ | N12 wx, N2 wy => f12 wx (extend2 9 wy)
+ | N12 wx, N3 wy => f12 wx (extend3 8 wy)
+ | N12 wx, N4 wy => f12 wx (extend4 7 wy)
+ | N12 wx, N5 wy => f12 wx (extend5 6 wy)
+ | N12 wx, N6 wy => f12 wx (extend6 5 wy)
+ | N12 wx, N7 wy => f12 wx (extend7 4 wy)
+ | N12 wx, N8 wy => f12 wx (extend8 3 wy)
+ | N12 wx, N9 wy => f12 wx (extend9 2 wy)
+ | N12 wx, N10 wy => f12 wx (extend10 1 wy)
+ | N12 wx, N11 wy => f12 wx (extend11 0 wy)
+ | N12 wx, N12 wy => f12 wx wy
+ | N12 wx, N13 wy => f13 (extend12 0 wx) wy
+ | N12 wx, Nn m wy => fnn m (extend13 m (extend12 0 wx)) wy
+ | N13 wx, N0 wy => f13 wx (extend0 12 wy)
+ | N13 wx, N1 wy => f13 wx (extend1 11 wy)
+ | N13 wx, N2 wy => f13 wx (extend2 10 wy)
+ | N13 wx, N3 wy => f13 wx (extend3 9 wy)
+ | N13 wx, N4 wy => f13 wx (extend4 8 wy)
+ | N13 wx, N5 wy => f13 wx (extend5 7 wy)
+ | N13 wx, N6 wy => f13 wx (extend6 6 wy)
+ | N13 wx, N7 wy => f13 wx (extend7 5 wy)
+ | N13 wx, N8 wy => f13 wx (extend8 4 wy)
+ | N13 wx, N9 wy => f13 wx (extend9 3 wy)
+ | N13 wx, N10 wy => f13 wx (extend10 2 wy)
+ | N13 wx, N11 wy => f13 wx (extend11 1 wy)
+ | N13 wx, N12 wy => f13 wx (extend12 0 wy)
+ | N13 wx, N13 wy => f13 wx wy
+ | N13 wx, Nn m wy => fnn m (extend13 m wx) wy
+ | Nn n wx, N0 wy => fnn n wx (extend13 n (extend0 12 wy))
+ | Nn n wx, N1 wy => fnn n wx (extend13 n (extend1 11 wy))
+ | Nn n wx, N2 wy => fnn n wx (extend13 n (extend2 10 wy))
+ | Nn n wx, N3 wy => fnn n wx (extend13 n (extend3 9 wy))
+ | Nn n wx, N4 wy => fnn n wx (extend13 n (extend4 8 wy))
+ | Nn n wx, N5 wy => fnn n wx (extend13 n (extend5 7 wy))
+ | Nn n wx, N6 wy => fnn n wx (extend13 n (extend6 6 wy))
+ | Nn n wx, N7 wy => fnn n wx (extend13 n (extend7 5 wy))
+ | Nn n wx, N8 wy => fnn n wx (extend13 n (extend8 4 wy))
+ | Nn n wx, N9 wy => fnn n wx (extend13 n (extend9 3 wy))
+ | Nn n wx, N10 wy => fnn n wx (extend13 n (extend10 2 wy))
+ | Nn n wx, N11 wy => fnn n wx (extend13 n (extend11 1 wy))
+ | Nn n wx, N12 wy => fnn n wx (extend13 n (extend12 0 wy))
+ | Nn n wx, N13 wy => fnn n wx (extend13 n wy)
+ | Nn n wx, Nn m wy =>
+ let mn := Max.max n m in
+ let d := diff n m in
+ fnn mn
+ (castm (diff_r n m) (extend_tr wx (snd d)))
+ (castm (diff_l n m) (extend_tr wy (fst d)))
+ end.
+
+ (* We level the two arguments before applying *)
+ (* the functions at each level (special zero case) *)
+ Definition same_level0 (x y: t_): res :=
+ Eval lazy zeta beta iota delta [extend0 extend1 extend2 extend3 extend4 extend5 extend6 extend7 extend8 extend9 extend10 extend11 extend12 extend13
+ GenBase.extend GenBase.extend_aux
+ ] in
+ match x with
+ | N0 wx =>
+ if w0_eq0 wx then f0t y else
+ match y with
+ | N0 wy => f0 wx wy
+ | N1 wy => f1 (extend0 0 wx) wy
+ | N2 wy => f2 (extend0 1 wx) wy
+ | N3 wy => f3 (extend0 2 wx) wy
+ | N4 wy => f4 (extend0 3 wx) wy
+ | N5 wy => f5 (extend0 4 wx) wy
+ | N6 wy => f6 (extend0 5 wx) wy
+ | N7 wy => f7 (extend0 6 wx) wy
+ | N8 wy => f8 (extend0 7 wx) wy
+ | N9 wy => f9 (extend0 8 wx) wy
+ | N10 wy => f10 (extend0 9 wx) wy
+ | N11 wy => f11 (extend0 10 wx) wy
+ | N12 wy => f12 (extend0 11 wx) wy
+ | N13 wy => f13 (extend0 12 wx) wy
+ | Nn m wy => fnn m (extend13 m (extend0 12 wx)) wy
+ end
+ | N1 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ f1 wx (extend0 0 wy)
+ | N1 wy => f1 wx wy
+ | N2 wy => f2 (extend1 0 wx) wy
+ | N3 wy => f3 (extend1 1 wx) wy
+ | N4 wy => f4 (extend1 2 wx) wy
+ | N5 wy => f5 (extend1 3 wx) wy
+ | N6 wy => f6 (extend1 4 wx) wy
+ | N7 wy => f7 (extend1 5 wx) wy
+ | N8 wy => f8 (extend1 6 wx) wy
+ | N9 wy => f9 (extend1 7 wx) wy
+ | N10 wy => f10 (extend1 8 wx) wy
+ | N11 wy => f11 (extend1 9 wx) wy
+ | N12 wy => f12 (extend1 10 wx) wy
+ | N13 wy => f13 (extend1 11 wx) wy
+ | Nn m wy => fnn m (extend13 m (extend1 11 wx)) wy
+ end
+ | N2 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ f2 wx (extend0 1 wy)
+ | N1 wy =>
+ f2 wx (extend1 0 wy)
+ | N2 wy => f2 wx wy
+ | N3 wy => f3 (extend2 0 wx) wy
+ | N4 wy => f4 (extend2 1 wx) wy
+ | N5 wy => f5 (extend2 2 wx) wy
+ | N6 wy => f6 (extend2 3 wx) wy
+ | N7 wy => f7 (extend2 4 wx) wy
+ | N8 wy => f8 (extend2 5 wx) wy
+ | N9 wy => f9 (extend2 6 wx) wy
+ | N10 wy => f10 (extend2 7 wx) wy
+ | N11 wy => f11 (extend2 8 wx) wy
+ | N12 wy => f12 (extend2 9 wx) wy
+ | N13 wy => f13 (extend2 10 wx) wy
+ | Nn m wy => fnn m (extend13 m (extend2 10 wx)) wy
+ end
+ | N3 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ f3 wx (extend0 2 wy)
+ | N1 wy =>
+ f3 wx (extend1 1 wy)
+ | N2 wy =>
+ f3 wx (extend2 0 wy)
+ | N3 wy => f3 wx wy
+ | N4 wy => f4 (extend3 0 wx) wy
+ | N5 wy => f5 (extend3 1 wx) wy
+ | N6 wy => f6 (extend3 2 wx) wy
+ | N7 wy => f7 (extend3 3 wx) wy
+ | N8 wy => f8 (extend3 4 wx) wy
+ | N9 wy => f9 (extend3 5 wx) wy
+ | N10 wy => f10 (extend3 6 wx) wy
+ | N11 wy => f11 (extend3 7 wx) wy
+ | N12 wy => f12 (extend3 8 wx) wy
+ | N13 wy => f13 (extend3 9 wx) wy
+ | Nn m wy => fnn m (extend13 m (extend3 9 wx)) wy
+ end
+ | N4 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ f4 wx (extend0 3 wy)
+ | N1 wy =>
+ f4 wx (extend1 2 wy)
+ | N2 wy =>
+ f4 wx (extend2 1 wy)
+ | N3 wy =>
+ f4 wx (extend3 0 wy)
+ | N4 wy => f4 wx wy
+ | N5 wy => f5 (extend4 0 wx) wy
+ | N6 wy => f6 (extend4 1 wx) wy
+ | N7 wy => f7 (extend4 2 wx) wy
+ | N8 wy => f8 (extend4 3 wx) wy
+ | N9 wy => f9 (extend4 4 wx) wy
+ | N10 wy => f10 (extend4 5 wx) wy
+ | N11 wy => f11 (extend4 6 wx) wy
+ | N12 wy => f12 (extend4 7 wx) wy
+ | N13 wy => f13 (extend4 8 wx) wy
+ | Nn m wy => fnn m (extend13 m (extend4 8 wx)) wy
+ end
+ | N5 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ f5 wx (extend0 4 wy)
+ | N1 wy =>
+ f5 wx (extend1 3 wy)
+ | N2 wy =>
+ f5 wx (extend2 2 wy)
+ | N3 wy =>
+ f5 wx (extend3 1 wy)
+ | N4 wy =>
+ f5 wx (extend4 0 wy)
+ | N5 wy => f5 wx wy
+ | N6 wy => f6 (extend5 0 wx) wy
+ | N7 wy => f7 (extend5 1 wx) wy
+ | N8 wy => f8 (extend5 2 wx) wy
+ | N9 wy => f9 (extend5 3 wx) wy
+ | N10 wy => f10 (extend5 4 wx) wy
+ | N11 wy => f11 (extend5 5 wx) wy
+ | N12 wy => f12 (extend5 6 wx) wy
+ | N13 wy => f13 (extend5 7 wx) wy
+ | Nn m wy => fnn m (extend13 m (extend5 7 wx)) wy
+ end
+ | N6 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ f6 wx (extend0 5 wy)
+ | N1 wy =>
+ f6 wx (extend1 4 wy)
+ | N2 wy =>
+ f6 wx (extend2 3 wy)
+ | N3 wy =>
+ f6 wx (extend3 2 wy)
+ | N4 wy =>
+ f6 wx (extend4 1 wy)
+ | N5 wy =>
+ f6 wx (extend5 0 wy)
+ | N6 wy => f6 wx wy
+ | N7 wy => f7 (extend6 0 wx) wy
+ | N8 wy => f8 (extend6 1 wx) wy
+ | N9 wy => f9 (extend6 2 wx) wy
+ | N10 wy => f10 (extend6 3 wx) wy
+ | N11 wy => f11 (extend6 4 wx) wy
+ | N12 wy => f12 (extend6 5 wx) wy
+ | N13 wy => f13 (extend6 6 wx) wy
+ | Nn m wy => fnn m (extend13 m (extend6 6 wx)) wy
+ end
+ | N7 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ f7 wx (extend0 6 wy)
+ | N1 wy =>
+ f7 wx (extend1 5 wy)
+ | N2 wy =>
+ f7 wx (extend2 4 wy)
+ | N3 wy =>
+ f7 wx (extend3 3 wy)
+ | N4 wy =>
+ f7 wx (extend4 2 wy)
+ | N5 wy =>
+ f7 wx (extend5 1 wy)
+ | N6 wy =>
+ f7 wx (extend6 0 wy)
+ | N7 wy => f7 wx wy
+ | N8 wy => f8 (extend7 0 wx) wy
+ | N9 wy => f9 (extend7 1 wx) wy
+ | N10 wy => f10 (extend7 2 wx) wy
+ | N11 wy => f11 (extend7 3 wx) wy
+ | N12 wy => f12 (extend7 4 wx) wy
+ | N13 wy => f13 (extend7 5 wx) wy
+ | Nn m wy => fnn m (extend13 m (extend7 5 wx)) wy
+ end
+ | N8 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ f8 wx (extend0 7 wy)
+ | N1 wy =>
+ f8 wx (extend1 6 wy)
+ | N2 wy =>
+ f8 wx (extend2 5 wy)
+ | N3 wy =>
+ f8 wx (extend3 4 wy)
+ | N4 wy =>
+ f8 wx (extend4 3 wy)
+ | N5 wy =>
+ f8 wx (extend5 2 wy)
+ | N6 wy =>
+ f8 wx (extend6 1 wy)
+ | N7 wy =>
+ f8 wx (extend7 0 wy)
+ | N8 wy => f8 wx wy
+ | N9 wy => f9 (extend8 0 wx) wy
+ | N10 wy => f10 (extend8 1 wx) wy
+ | N11 wy => f11 (extend8 2 wx) wy
+ | N12 wy => f12 (extend8 3 wx) wy
+ | N13 wy => f13 (extend8 4 wx) wy
+ | Nn m wy => fnn m (extend13 m (extend8 4 wx)) wy
+ end
+ | N9 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ f9 wx (extend0 8 wy)
+ | N1 wy =>
+ f9 wx (extend1 7 wy)
+ | N2 wy =>
+ f9 wx (extend2 6 wy)
+ | N3 wy =>
+ f9 wx (extend3 5 wy)
+ | N4 wy =>
+ f9 wx (extend4 4 wy)
+ | N5 wy =>
+ f9 wx (extend5 3 wy)
+ | N6 wy =>
+ f9 wx (extend6 2 wy)
+ | N7 wy =>
+ f9 wx (extend7 1 wy)
+ | N8 wy =>
+ f9 wx (extend8 0 wy)
+ | N9 wy => f9 wx wy
+ | N10 wy => f10 (extend9 0 wx) wy
+ | N11 wy => f11 (extend9 1 wx) wy
+ | N12 wy => f12 (extend9 2 wx) wy
+ | N13 wy => f13 (extend9 3 wx) wy
+ | Nn m wy => fnn m (extend13 m (extend9 3 wx)) wy
+ end
+ | N10 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ f10 wx (extend0 9 wy)
+ | N1 wy =>
+ f10 wx (extend1 8 wy)
+ | N2 wy =>
+ f10 wx (extend2 7 wy)
+ | N3 wy =>
+ f10 wx (extend3 6 wy)
+ | N4 wy =>
+ f10 wx (extend4 5 wy)
+ | N5 wy =>
+ f10 wx (extend5 4 wy)
+ | N6 wy =>
+ f10 wx (extend6 3 wy)
+ | N7 wy =>
+ f10 wx (extend7 2 wy)
+ | N8 wy =>
+ f10 wx (extend8 1 wy)
+ | N9 wy =>
+ f10 wx (extend9 0 wy)
+ | N10 wy => f10 wx wy
+ | N11 wy => f11 (extend10 0 wx) wy
+ | N12 wy => f12 (extend10 1 wx) wy
+ | N13 wy => f13 (extend10 2 wx) wy
+ | Nn m wy => fnn m (extend13 m (extend10 2 wx)) wy
+ end
+ | N11 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ f11 wx (extend0 10 wy)
+ | N1 wy =>
+ f11 wx (extend1 9 wy)
+ | N2 wy =>
+ f11 wx (extend2 8 wy)
+ | N3 wy =>
+ f11 wx (extend3 7 wy)
+ | N4 wy =>
+ f11 wx (extend4 6 wy)
+ | N5 wy =>
+ f11 wx (extend5 5 wy)
+ | N6 wy =>
+ f11 wx (extend6 4 wy)
+ | N7 wy =>
+ f11 wx (extend7 3 wy)
+ | N8 wy =>
+ f11 wx (extend8 2 wy)
+ | N9 wy =>
+ f11 wx (extend9 1 wy)
+ | N10 wy =>
+ f11 wx (extend10 0 wy)
+ | N11 wy => f11 wx wy
+ | N12 wy => f12 (extend11 0 wx) wy
+ | N13 wy => f13 (extend11 1 wx) wy
+ | Nn m wy => fnn m (extend13 m (extend11 1 wx)) wy
+ end
+ | N12 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ f12 wx (extend0 11 wy)
+ | N1 wy =>
+ f12 wx (extend1 10 wy)
+ | N2 wy =>
+ f12 wx (extend2 9 wy)
+ | N3 wy =>
+ f12 wx (extend3 8 wy)
+ | N4 wy =>
+ f12 wx (extend4 7 wy)
+ | N5 wy =>
+ f12 wx (extend5 6 wy)
+ | N6 wy =>
+ f12 wx (extend6 5 wy)
+ | N7 wy =>
+ f12 wx (extend7 4 wy)
+ | N8 wy =>
+ f12 wx (extend8 3 wy)
+ | N9 wy =>
+ f12 wx (extend9 2 wy)
+ | N10 wy =>
+ f12 wx (extend10 1 wy)
+ | N11 wy =>
+ f12 wx (extend11 0 wy)
+ | N12 wy => f12 wx wy
+ | N13 wy => f13 (extend12 0 wx) wy
+ | Nn m wy => fnn m (extend13 m (extend12 0 wx)) wy
+ end
+ | N13 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ f13 wx (extend0 12 wy)
+ | N1 wy =>
+ f13 wx (extend1 11 wy)
+ | N2 wy =>
+ f13 wx (extend2 10 wy)
+ | N3 wy =>
+ f13 wx (extend3 9 wy)
+ | N4 wy =>
+ f13 wx (extend4 8 wy)
+ | N5 wy =>
+ f13 wx (extend5 7 wy)
+ | N6 wy =>
+ f13 wx (extend6 6 wy)
+ | N7 wy =>
+ f13 wx (extend7 5 wy)
+ | N8 wy =>
+ f13 wx (extend8 4 wy)
+ | N9 wy =>
+ f13 wx (extend9 3 wy)
+ | N10 wy =>
+ f13 wx (extend10 2 wy)
+ | N11 wy =>
+ f13 wx (extend11 1 wy)
+ | N12 wy =>
+ f13 wx (extend12 0 wy)
+ | N13 wy => f13 wx wy
+ | Nn m wy => fnn m (extend13 m wx) wy
+ end
+ | Nn n wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ fnn n wx (extend13 n (extend0 12 wy))
+ | N1 wy =>
+ fnn n wx (extend13 n (extend1 11 wy))
+ | N2 wy =>
+ fnn n wx (extend13 n (extend2 10 wy))
+ | N3 wy =>
+ fnn n wx (extend13 n (extend3 9 wy))
+ | N4 wy =>
+ fnn n wx (extend13 n (extend4 8 wy))
+ | N5 wy =>
+ fnn n wx (extend13 n (extend5 7 wy))
+ | N6 wy =>
+ fnn n wx (extend13 n (extend6 6 wy))
+ | N7 wy =>
+ fnn n wx (extend13 n (extend7 5 wy))
+ | N8 wy =>
+ fnn n wx (extend13 n (extend8 4 wy))
+ | N9 wy =>
+ fnn n wx (extend13 n (extend9 3 wy))
+ | N10 wy =>
+ fnn n wx (extend13 n (extend10 2 wy))
+ | N11 wy =>
+ fnn n wx (extend13 n (extend11 1 wy))
+ | N12 wy =>
+ fnn n wx (extend13 n (extend12 0 wy))
+ | N13 wy =>
+ fnn n wx (extend13 n wy)
+ | Nn m wy =>
+ let mn := Max.max n m in
+ let d := diff n m in
+ fnn mn
+ (castm (diff_r n m) (extend_tr wx (snd d)))
+ (castm (diff_l n m) (extend_tr wy (fst d)))
+ end
+ end.
+
+ (* We iter the smaller argument with the bigger *)
+ Definition iter (x y: t_): res :=
+ Eval lazy zeta beta iota delta [extend0 extend1 extend2 extend3 extend4 extend5 extend6 extend7 extend8 extend9 extend10 extend11 extend12 extend13
+ GenBase.extend GenBase.extend_aux
+ ] in
+ match x, y with
+ | N0 wx, N0 wy => f0 wx wy
+ | N0 wx, N1 wy => f0n 0 wx wy
+ | N0 wx, N2 wy => f0n 1 wx wy
+ | N0 wx, N3 wy => f0n 2 wx wy
+ | N0 wx, N4 wy => f0n 3 wx wy
+ | N0 wx, N5 wy => f0n 4 wx wy
+ | N0 wx, N6 wy => f0n 5 wx wy
+ | N0 wx, N7 wy => f0n 6 wx wy
+ | N0 wx, N8 wy => f0n 7 wx wy
+ | N0 wx, N9 wy => f0n 8 wx wy
+ | N0 wx, N10 wy => f0n 9 wx wy
+ | N0 wx, N11 wy => f0n 10 wx wy
+ | N0 wx, N12 wy => f0n 11 wx wy
+ | N0 wx, N13 wy => f0n 12 wx wy
+ | N0 wx, Nn m wy => f13n m (extend0 12 wx) wy
+ | N1 wx, N0 wy => fn0 0 wx wy
+ | N1 wx, N1 wy => f1 wx wy
+ | N1 wx, N2 wy => f1n 0 wx wy
+ | N1 wx, N3 wy => f1n 1 wx wy
+ | N1 wx, N4 wy => f1n 2 wx wy
+ | N1 wx, N5 wy => f1n 3 wx wy
+ | N1 wx, N6 wy => f1n 4 wx wy
+ | N1 wx, N7 wy => f1n 5 wx wy
+ | N1 wx, N8 wy => f1n 6 wx wy
+ | N1 wx, N9 wy => f1n 7 wx wy
+ | N1 wx, N10 wy => f1n 8 wx wy
+ | N1 wx, N11 wy => f1n 9 wx wy
+ | N1 wx, N12 wy => f1n 10 wx wy
+ | N1 wx, N13 wy => f1n 11 wx wy
+ | N1 wx, Nn m wy => f13n m (extend1 11 wx) wy
+ | N2 wx, N0 wy => fn0 1 wx wy
+ | N2 wx, N1 wy => fn1 0 wx wy
+ | N2 wx, N2 wy => f2 wx wy
+ | N2 wx, N3 wy => f2n 0 wx wy
+ | N2 wx, N4 wy => f2n 1 wx wy
+ | N2 wx, N5 wy => f2n 2 wx wy
+ | N2 wx, N6 wy => f2n 3 wx wy
+ | N2 wx, N7 wy => f2n 4 wx wy
+ | N2 wx, N8 wy => f2n 5 wx wy
+ | N2 wx, N9 wy => f2n 6 wx wy
+ | N2 wx, N10 wy => f2n 7 wx wy
+ | N2 wx, N11 wy => f2n 8 wx wy
+ | N2 wx, N12 wy => f2n 9 wx wy
+ | N2 wx, N13 wy => f2n 10 wx wy
+ | N2 wx, Nn m wy => f13n m (extend2 10 wx) wy
+ | N3 wx, N0 wy => fn0 2 wx wy
+ | N3 wx, N1 wy => fn1 1 wx wy
+ | N3 wx, N2 wy => fn2 0 wx wy
+ | N3 wx, N3 wy => f3 wx wy
+ | N3 wx, N4 wy => f3n 0 wx wy
+ | N3 wx, N5 wy => f3n 1 wx wy
+ | N3 wx, N6 wy => f3n 2 wx wy
+ | N3 wx, N7 wy => f3n 3 wx wy
+ | N3 wx, N8 wy => f3n 4 wx wy
+ | N3 wx, N9 wy => f3n 5 wx wy
+ | N3 wx, N10 wy => f3n 6 wx wy
+ | N3 wx, N11 wy => f3n 7 wx wy
+ | N3 wx, N12 wy => f3n 8 wx wy
+ | N3 wx, N13 wy => f3n 9 wx wy
+ | N3 wx, Nn m wy => f13n m (extend3 9 wx) wy
+ | N4 wx, N0 wy => fn0 3 wx wy
+ | N4 wx, N1 wy => fn1 2 wx wy
+ | N4 wx, N2 wy => fn2 1 wx wy
+ | N4 wx, N3 wy => fn3 0 wx wy
+ | N4 wx, N4 wy => f4 wx wy
+ | N4 wx, N5 wy => f4n 0 wx wy
+ | N4 wx, N6 wy => f4n 1 wx wy
+ | N4 wx, N7 wy => f4n 2 wx wy
+ | N4 wx, N8 wy => f4n 3 wx wy
+ | N4 wx, N9 wy => f4n 4 wx wy
+ | N4 wx, N10 wy => f4n 5 wx wy
+ | N4 wx, N11 wy => f4n 6 wx wy
+ | N4 wx, N12 wy => f4n 7 wx wy
+ | N4 wx, N13 wy => f4n 8 wx wy
+ | N4 wx, Nn m wy => f13n m (extend4 8 wx) wy
+ | N5 wx, N0 wy => fn0 4 wx wy
+ | N5 wx, N1 wy => fn1 3 wx wy
+ | N5 wx, N2 wy => fn2 2 wx wy
+ | N5 wx, N3 wy => fn3 1 wx wy
+ | N5 wx, N4 wy => fn4 0 wx wy
+ | N5 wx, N5 wy => f5 wx wy
+ | N5 wx, N6 wy => f5n 0 wx wy
+ | N5 wx, N7 wy => f5n 1 wx wy
+ | N5 wx, N8 wy => f5n 2 wx wy
+ | N5 wx, N9 wy => f5n 3 wx wy
+ | N5 wx, N10 wy => f5n 4 wx wy
+ | N5 wx, N11 wy => f5n 5 wx wy
+ | N5 wx, N12 wy => f5n 6 wx wy
+ | N5 wx, N13 wy => f5n 7 wx wy
+ | N5 wx, Nn m wy => f13n m (extend5 7 wx) wy
+ | N6 wx, N0 wy => fn0 5 wx wy
+ | N6 wx, N1 wy => fn1 4 wx wy
+ | N6 wx, N2 wy => fn2 3 wx wy
+ | N6 wx, N3 wy => fn3 2 wx wy
+ | N6 wx, N4 wy => fn4 1 wx wy
+ | N6 wx, N5 wy => fn5 0 wx wy
+ | N6 wx, N6 wy => f6 wx wy
+ | N6 wx, N7 wy => f6n 0 wx wy
+ | N6 wx, N8 wy => f6n 1 wx wy
+ | N6 wx, N9 wy => f6n 2 wx wy
+ | N6 wx, N10 wy => f6n 3 wx wy
+ | N6 wx, N11 wy => f6n 4 wx wy
+ | N6 wx, N12 wy => f6n 5 wx wy
+ | N6 wx, N13 wy => f6n 6 wx wy
+ | N6 wx, Nn m wy => f13n m (extend6 6 wx) wy
+ | N7 wx, N0 wy => fn0 6 wx wy
+ | N7 wx, N1 wy => fn1 5 wx wy
+ | N7 wx, N2 wy => fn2 4 wx wy
+ | N7 wx, N3 wy => fn3 3 wx wy
+ | N7 wx, N4 wy => fn4 2 wx wy
+ | N7 wx, N5 wy => fn5 1 wx wy
+ | N7 wx, N6 wy => fn6 0 wx wy
+ | N7 wx, N7 wy => f7 wx wy
+ | N7 wx, N8 wy => f7n 0 wx wy
+ | N7 wx, N9 wy => f7n 1 wx wy
+ | N7 wx, N10 wy => f7n 2 wx wy
+ | N7 wx, N11 wy => f7n 3 wx wy
+ | N7 wx, N12 wy => f7n 4 wx wy
+ | N7 wx, N13 wy => f7n 5 wx wy
+ | N7 wx, Nn m wy => f13n m (extend7 5 wx) wy
+ | N8 wx, N0 wy => fn0 7 wx wy
+ | N8 wx, N1 wy => fn1 6 wx wy
+ | N8 wx, N2 wy => fn2 5 wx wy
+ | N8 wx, N3 wy => fn3 4 wx wy
+ | N8 wx, N4 wy => fn4 3 wx wy
+ | N8 wx, N5 wy => fn5 2 wx wy
+ | N8 wx, N6 wy => fn6 1 wx wy
+ | N8 wx, N7 wy => fn7 0 wx wy
+ | N8 wx, N8 wy => f8 wx wy
+ | N8 wx, N9 wy => f8n 0 wx wy
+ | N8 wx, N10 wy => f8n 1 wx wy
+ | N8 wx, N11 wy => f8n 2 wx wy
+ | N8 wx, N12 wy => f8n 3 wx wy
+ | N8 wx, N13 wy => f8n 4 wx wy
+ | N8 wx, Nn m wy => f13n m (extend8 4 wx) wy
+ | N9 wx, N0 wy => fn0 8 wx wy
+ | N9 wx, N1 wy => fn1 7 wx wy
+ | N9 wx, N2 wy => fn2 6 wx wy
+ | N9 wx, N3 wy => fn3 5 wx wy
+ | N9 wx, N4 wy => fn4 4 wx wy
+ | N9 wx, N5 wy => fn5 3 wx wy
+ | N9 wx, N6 wy => fn6 2 wx wy
+ | N9 wx, N7 wy => fn7 1 wx wy
+ | N9 wx, N8 wy => fn8 0 wx wy
+ | N9 wx, N9 wy => f9 wx wy
+ | N9 wx, N10 wy => f9n 0 wx wy
+ | N9 wx, N11 wy => f9n 1 wx wy
+ | N9 wx, N12 wy => f9n 2 wx wy
+ | N9 wx, N13 wy => f9n 3 wx wy
+ | N9 wx, Nn m wy => f13n m (extend9 3 wx) wy
+ | N10 wx, N0 wy => fn0 9 wx wy
+ | N10 wx, N1 wy => fn1 8 wx wy
+ | N10 wx, N2 wy => fn2 7 wx wy
+ | N10 wx, N3 wy => fn3 6 wx wy
+ | N10 wx, N4 wy => fn4 5 wx wy
+ | N10 wx, N5 wy => fn5 4 wx wy
+ | N10 wx, N6 wy => fn6 3 wx wy
+ | N10 wx, N7 wy => fn7 2 wx wy
+ | N10 wx, N8 wy => fn8 1 wx wy
+ | N10 wx, N9 wy => fn9 0 wx wy
+ | N10 wx, N10 wy => f10 wx wy
+ | N10 wx, N11 wy => f10n 0 wx wy
+ | N10 wx, N12 wy => f10n 1 wx wy
+ | N10 wx, N13 wy => f10n 2 wx wy
+ | N10 wx, Nn m wy => f13n m (extend10 2 wx) wy
+ | N11 wx, N0 wy => fn0 10 wx wy
+ | N11 wx, N1 wy => fn1 9 wx wy
+ | N11 wx, N2 wy => fn2 8 wx wy
+ | N11 wx, N3 wy => fn3 7 wx wy
+ | N11 wx, N4 wy => fn4 6 wx wy
+ | N11 wx, N5 wy => fn5 5 wx wy
+ | N11 wx, N6 wy => fn6 4 wx wy
+ | N11 wx, N7 wy => fn7 3 wx wy
+ | N11 wx, N8 wy => fn8 2 wx wy
+ | N11 wx, N9 wy => fn9 1 wx wy
+ | N11 wx, N10 wy => fn10 0 wx wy
+ | N11 wx, N11 wy => f11 wx wy
+ | N11 wx, N12 wy => f11n 0 wx wy
+ | N11 wx, N13 wy => f11n 1 wx wy
+ | N11 wx, Nn m wy => f13n m (extend11 1 wx) wy
+ | N12 wx, N0 wy => fn0 11 wx wy
+ | N12 wx, N1 wy => fn1 10 wx wy
+ | N12 wx, N2 wy => fn2 9 wx wy
+ | N12 wx, N3 wy => fn3 8 wx wy
+ | N12 wx, N4 wy => fn4 7 wx wy
+ | N12 wx, N5 wy => fn5 6 wx wy
+ | N12 wx, N6 wy => fn6 5 wx wy
+ | N12 wx, N7 wy => fn7 4 wx wy
+ | N12 wx, N8 wy => fn8 3 wx wy
+ | N12 wx, N9 wy => fn9 2 wx wy
+ | N12 wx, N10 wy => fn10 1 wx wy
+ | N12 wx, N11 wy => fn11 0 wx wy
+ | N12 wx, N12 wy => f12 wx wy
+ | N12 wx, N13 wy => f12n 0 wx wy
+ | N12 wx, Nn m wy => f13n m (extend12 0 wx) wy
+ | N13 wx, N0 wy => fn0 12 wx wy
+ | N13 wx, N1 wy => fn1 11 wx wy
+ | N13 wx, N2 wy => fn2 10 wx wy
+ | N13 wx, N3 wy => fn3 9 wx wy
+ | N13 wx, N4 wy => fn4 8 wx wy
+ | N13 wx, N5 wy => fn5 7 wx wy
+ | N13 wx, N6 wy => fn6 6 wx wy
+ | N13 wx, N7 wy => fn7 5 wx wy
+ | N13 wx, N8 wy => fn8 4 wx wy
+ | N13 wx, N9 wy => fn9 3 wx wy
+ | N13 wx, N10 wy => fn10 2 wx wy
+ | N13 wx, N11 wy => fn11 1 wx wy
+ | N13 wx, N12 wy => fn12 0 wx wy
+ | N13 wx, N13 wy => f13 wx wy
+ | N13 wx, Nn m wy => f13n m wx wy
+ | Nn n wx, N0 wy => fn13 n wx (extend0 12 wy)
+ | Nn n wx, N1 wy => fn13 n wx (extend1 11 wy)
+ | Nn n wx, N2 wy => fn13 n wx (extend2 10 wy)
+ | Nn n wx, N3 wy => fn13 n wx (extend3 9 wy)
+ | Nn n wx, N4 wy => fn13 n wx (extend4 8 wy)
+ | Nn n wx, N5 wy => fn13 n wx (extend5 7 wy)
+ | Nn n wx, N6 wy => fn13 n wx (extend6 6 wy)
+ | Nn n wx, N7 wy => fn13 n wx (extend7 5 wy)
+ | Nn n wx, N8 wy => fn13 n wx (extend8 4 wy)
+ | Nn n wx, N9 wy => fn13 n wx (extend9 3 wy)
+ | Nn n wx, N10 wy => fn13 n wx (extend10 2 wy)
+ | Nn n wx, N11 wy => fn13 n wx (extend11 1 wy)
+ | Nn n wx, N12 wy => fn13 n wx (extend12 0 wy)
+ | Nn n wx, N13 wy => fn13 n wx wy
+ | Nn n wx, Nn m wy => fnm n m wx wy
+ end.
+
+ (* We iter the smaller argument with the bigger (zero case) *)
+ Definition iter0 (x y: t_): res :=
+ Eval lazy zeta beta iota delta [extend0 extend1 extend2 extend3 extend4 extend5 extend6 extend7 extend8 extend9 extend10 extend11 extend12 extend13
+ GenBase.extend GenBase.extend_aux
+ ] in
+ match x with
+ | N0 wx =>
+ if w0_eq0 wx then f0t y else
+ match y with
+ | N0 wy => f0 wx wy
+ | N1 wy => f0n 0 wx wy
+ | N2 wy => f0n 1 wx wy
+ | N3 wy => f0n 2 wx wy
+ | N4 wy => f0n 3 wx wy
+ | N5 wy => f0n 4 wx wy
+ | N6 wy => f0n 5 wx wy
+ | N7 wy => f0n 6 wx wy
+ | N8 wy => f0n 7 wx wy
+ | N9 wy => f0n 8 wx wy
+ | N10 wy => f0n 9 wx wy
+ | N11 wy => f0n 10 wx wy
+ | N12 wy => f0n 11 wx wy
+ | N13 wy => f0n 12 wx wy
+ | Nn m wy => f13n m (extend0 12 wx) wy
+ end
+ | N1 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ fn0 0 wx wy
+ | N1 wy => f1 wx wy
+ | N2 wy => f1n 0 wx wy
+ | N3 wy => f1n 1 wx wy
+ | N4 wy => f1n 2 wx wy
+ | N5 wy => f1n 3 wx wy
+ | N6 wy => f1n 4 wx wy
+ | N7 wy => f1n 5 wx wy
+ | N8 wy => f1n 6 wx wy
+ | N9 wy => f1n 7 wx wy
+ | N10 wy => f1n 8 wx wy
+ | N11 wy => f1n 9 wx wy
+ | N12 wy => f1n 10 wx wy
+ | N13 wy => f1n 11 wx wy
+ | Nn m wy => f13n m (extend1 11 wx) wy
+ end
+ | N2 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ fn0 1 wx wy
+ | N1 wy =>
+ fn1 0 wx wy
+ | N2 wy => f2 wx wy
+ | N3 wy => f2n 0 wx wy
+ | N4 wy => f2n 1 wx wy
+ | N5 wy => f2n 2 wx wy
+ | N6 wy => f2n 3 wx wy
+ | N7 wy => f2n 4 wx wy
+ | N8 wy => f2n 5 wx wy
+ | N9 wy => f2n 6 wx wy
+ | N10 wy => f2n 7 wx wy
+ | N11 wy => f2n 8 wx wy
+ | N12 wy => f2n 9 wx wy
+ | N13 wy => f2n 10 wx wy
+ | Nn m wy => f13n m (extend2 10 wx) wy
+ end
+ | N3 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ fn0 2 wx wy
+ | N1 wy =>
+ fn1 1 wx wy
+ | N2 wy =>
+ fn2 0 wx wy
+ | N3 wy => f3 wx wy
+ | N4 wy => f3n 0 wx wy
+ | N5 wy => f3n 1 wx wy
+ | N6 wy => f3n 2 wx wy
+ | N7 wy => f3n 3 wx wy
+ | N8 wy => f3n 4 wx wy
+ | N9 wy => f3n 5 wx wy
+ | N10 wy => f3n 6 wx wy
+ | N11 wy => f3n 7 wx wy
+ | N12 wy => f3n 8 wx wy
+ | N13 wy => f3n 9 wx wy
+ | Nn m wy => f13n m (extend3 9 wx) wy
+ end
+ | N4 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ fn0 3 wx wy
+ | N1 wy =>
+ fn1 2 wx wy
+ | N2 wy =>
+ fn2 1 wx wy
+ | N3 wy =>
+ fn3 0 wx wy
+ | N4 wy => f4 wx wy
+ | N5 wy => f4n 0 wx wy
+ | N6 wy => f4n 1 wx wy
+ | N7 wy => f4n 2 wx wy
+ | N8 wy => f4n 3 wx wy
+ | N9 wy => f4n 4 wx wy
+ | N10 wy => f4n 5 wx wy
+ | N11 wy => f4n 6 wx wy
+ | N12 wy => f4n 7 wx wy
+ | N13 wy => f4n 8 wx wy
+ | Nn m wy => f13n m (extend4 8 wx) wy
+ end
+ | N5 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ fn0 4 wx wy
+ | N1 wy =>
+ fn1 3 wx wy
+ | N2 wy =>
+ fn2 2 wx wy
+ | N3 wy =>
+ fn3 1 wx wy
+ | N4 wy =>
+ fn4 0 wx wy
+ | N5 wy => f5 wx wy
+ | N6 wy => f5n 0 wx wy
+ | N7 wy => f5n 1 wx wy
+ | N8 wy => f5n 2 wx wy
+ | N9 wy => f5n 3 wx wy
+ | N10 wy => f5n 4 wx wy
+ | N11 wy => f5n 5 wx wy
+ | N12 wy => f5n 6 wx wy
+ | N13 wy => f5n 7 wx wy
+ | Nn m wy => f13n m (extend5 7 wx) wy
+ end
+ | N6 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ fn0 5 wx wy
+ | N1 wy =>
+ fn1 4 wx wy
+ | N2 wy =>
+ fn2 3 wx wy
+ | N3 wy =>
+ fn3 2 wx wy
+ | N4 wy =>
+ fn4 1 wx wy
+ | N5 wy =>
+ fn5 0 wx wy
+ | N6 wy => f6 wx wy
+ | N7 wy => f6n 0 wx wy
+ | N8 wy => f6n 1 wx wy
+ | N9 wy => f6n 2 wx wy
+ | N10 wy => f6n 3 wx wy
+ | N11 wy => f6n 4 wx wy
+ | N12 wy => f6n 5 wx wy
+ | N13 wy => f6n 6 wx wy
+ | Nn m wy => f13n m (extend6 6 wx) wy
+ end
+ | N7 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ fn0 6 wx wy
+ | N1 wy =>
+ fn1 5 wx wy
+ | N2 wy =>
+ fn2 4 wx wy
+ | N3 wy =>
+ fn3 3 wx wy
+ | N4 wy =>
+ fn4 2 wx wy
+ | N5 wy =>
+ fn5 1 wx wy
+ | N6 wy =>
+ fn6 0 wx wy
+ | N7 wy => f7 wx wy
+ | N8 wy => f7n 0 wx wy
+ | N9 wy => f7n 1 wx wy
+ | N10 wy => f7n 2 wx wy
+ | N11 wy => f7n 3 wx wy
+ | N12 wy => f7n 4 wx wy
+ | N13 wy => f7n 5 wx wy
+ | Nn m wy => f13n m (extend7 5 wx) wy
+ end
+ | N8 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ fn0 7 wx wy
+ | N1 wy =>
+ fn1 6 wx wy
+ | N2 wy =>
+ fn2 5 wx wy
+ | N3 wy =>
+ fn3 4 wx wy
+ | N4 wy =>
+ fn4 3 wx wy
+ | N5 wy =>
+ fn5 2 wx wy
+ | N6 wy =>
+ fn6 1 wx wy
+ | N7 wy =>
+ fn7 0 wx wy
+ | N8 wy => f8 wx wy
+ | N9 wy => f8n 0 wx wy
+ | N10 wy => f8n 1 wx wy
+ | N11 wy => f8n 2 wx wy
+ | N12 wy => f8n 3 wx wy
+ | N13 wy => f8n 4 wx wy
+ | Nn m wy => f13n m (extend8 4 wx) wy
+ end
+ | N9 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ fn0 8 wx wy
+ | N1 wy =>
+ fn1 7 wx wy
+ | N2 wy =>
+ fn2 6 wx wy
+ | N3 wy =>
+ fn3 5 wx wy
+ | N4 wy =>
+ fn4 4 wx wy
+ | N5 wy =>
+ fn5 3 wx wy
+ | N6 wy =>
+ fn6 2 wx wy
+ | N7 wy =>
+ fn7 1 wx wy
+ | N8 wy =>
+ fn8 0 wx wy
+ | N9 wy => f9 wx wy
+ | N10 wy => f9n 0 wx wy
+ | N11 wy => f9n 1 wx wy
+ | N12 wy => f9n 2 wx wy
+ | N13 wy => f9n 3 wx wy
+ | Nn m wy => f13n m (extend9 3 wx) wy
+ end
+ | N10 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ fn0 9 wx wy
+ | N1 wy =>
+ fn1 8 wx wy
+ | N2 wy =>
+ fn2 7 wx wy
+ | N3 wy =>
+ fn3 6 wx wy
+ | N4 wy =>
+ fn4 5 wx wy
+ | N5 wy =>
+ fn5 4 wx wy
+ | N6 wy =>
+ fn6 3 wx wy
+ | N7 wy =>
+ fn7 2 wx wy
+ | N8 wy =>
+ fn8 1 wx wy
+ | N9 wy =>
+ fn9 0 wx wy
+ | N10 wy => f10 wx wy
+ | N11 wy => f10n 0 wx wy
+ | N12 wy => f10n 1 wx wy
+ | N13 wy => f10n 2 wx wy
+ | Nn m wy => f13n m (extend10 2 wx) wy
+ end
+ | N11 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ fn0 10 wx wy
+ | N1 wy =>
+ fn1 9 wx wy
+ | N2 wy =>
+ fn2 8 wx wy
+ | N3 wy =>
+ fn3 7 wx wy
+ | N4 wy =>
+ fn4 6 wx wy
+ | N5 wy =>
+ fn5 5 wx wy
+ | N6 wy =>
+ fn6 4 wx wy
+ | N7 wy =>
+ fn7 3 wx wy
+ | N8 wy =>
+ fn8 2 wx wy
+ | N9 wy =>
+ fn9 1 wx wy
+ | N10 wy =>
+ fn10 0 wx wy
+ | N11 wy => f11 wx wy
+ | N12 wy => f11n 0 wx wy
+ | N13 wy => f11n 1 wx wy
+ | Nn m wy => f13n m (extend11 1 wx) wy
+ end
+ | N12 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ fn0 11 wx wy
+ | N1 wy =>
+ fn1 10 wx wy
+ | N2 wy =>
+ fn2 9 wx wy
+ | N3 wy =>
+ fn3 8 wx wy
+ | N4 wy =>
+ fn4 7 wx wy
+ | N5 wy =>
+ fn5 6 wx wy
+ | N6 wy =>
+ fn6 5 wx wy
+ | N7 wy =>
+ fn7 4 wx wy
+ | N8 wy =>
+ fn8 3 wx wy
+ | N9 wy =>
+ fn9 2 wx wy
+ | N10 wy =>
+ fn10 1 wx wy
+ | N11 wy =>
+ fn11 0 wx wy
+ | N12 wy => f12 wx wy
+ | N13 wy => f12n 0 wx wy
+ | Nn m wy => f13n m (extend12 0 wx) wy
+ end
+ | N13 wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ fn0 12 wx wy
+ | N1 wy =>
+ fn1 11 wx wy
+ | N2 wy =>
+ fn2 10 wx wy
+ | N3 wy =>
+ fn3 9 wx wy
+ | N4 wy =>
+ fn4 8 wx wy
+ | N5 wy =>
+ fn5 7 wx wy
+ | N6 wy =>
+ fn6 6 wx wy
+ | N7 wy =>
+ fn7 5 wx wy
+ | N8 wy =>
+ fn8 4 wx wy
+ | N9 wy =>
+ fn9 3 wx wy
+ | N10 wy =>
+ fn10 2 wx wy
+ | N11 wy =>
+ fn11 1 wx wy
+ | N12 wy =>
+ fn12 0 wx wy
+ | N13 wy => f13 wx wy
+ | Nn m wy => f13n m wx wy
+ end
+ | Nn n wx =>
+ match y with
+ | N0 wy =>
+ if w0_eq0 wy then ft0 x else
+ fn13 n wx (extend0 12 wy)
+ | N1 wy =>
+ fn13 n wx (extend1 11 wy)
+ | N2 wy =>
+ fn13 n wx (extend2 10 wy)
+ | N3 wy =>
+ fn13 n wx (extend3 9 wy)
+ | N4 wy =>
+ fn13 n wx (extend4 8 wy)
+ | N5 wy =>
+ fn13 n wx (extend5 7 wy)
+ | N6 wy =>
+ fn13 n wx (extend6 6 wy)
+ | N7 wy =>
+ fn13 n wx (extend7 5 wy)
+ | N8 wy =>
+ fn13 n wx (extend8 4 wy)
+ | N9 wy =>
+ fn13 n wx (extend9 3 wy)
+ | N10 wy =>
+ fn13 n wx (extend10 2 wy)
+ | N11 wy =>
+ fn13 n wx (extend11 1 wy)
+ | N12 wy =>
+ fn13 n wx (extend12 0 wy)
+ | N13 wy =>
+ fn13 n wx wy
+ | Nn m wy => fnm n m wx wy
+ end
+ end.
+
+ End LevelAndIter.
+
+ (***************************************************************)
+ (* *)
+ (* Reduction *)
+ (* *)
+ (***************************************************************)
+
+ Definition reduce_0 (x:w) := N0 x.
+ Definition reduce_1 :=
+ Eval lazy beta iota delta[reduce_n1] in
+ reduce_n1 _ _ zero w0_eq0 N0 N1.
+ Definition reduce_2 :=
+ Eval lazy beta iota delta[reduce_n1] in
+ reduce_n1 _ _ zero w1_eq0 reduce_1 N2.
+ Definition reduce_3 :=
+ Eval lazy beta iota delta[reduce_n1] in
+ reduce_n1 _ _ zero w2_eq0 reduce_2 N3.
+ Definition reduce_4 :=
+ Eval lazy beta iota delta[reduce_n1] in
+ reduce_n1 _ _ zero w3_eq0 reduce_3 N4.
+ Definition reduce_5 :=
+ Eval lazy beta iota delta[reduce_n1] in
+ reduce_n1 _ _ zero w4_eq0 reduce_4 N5.
+ Definition reduce_6 :=
+ Eval lazy beta iota delta[reduce_n1] in
+ reduce_n1 _ _ zero w5_eq0 reduce_5 N6.
+ Definition reduce_7 :=
+ Eval lazy beta iota delta[reduce_n1] in
+ reduce_n1 _ _ zero w6_eq0 reduce_6 N7.
+ Definition reduce_8 :=
+ Eval lazy beta iota delta[reduce_n1] in
+ reduce_n1 _ _ zero w7_eq0 reduce_7 N8.
+ Definition reduce_9 :=
+ Eval lazy beta iota delta[reduce_n1] in
+ reduce_n1 _ _ zero w8_eq0 reduce_8 N9.
+ Definition reduce_10 :=
+ Eval lazy beta iota delta[reduce_n1] in
+ reduce_n1 _ _ zero w9_eq0 reduce_9 N10.
+ Definition reduce_11 :=
+ Eval lazy beta iota delta[reduce_n1] in
+ reduce_n1 _ _ zero w10_eq0 reduce_10 N11.
+ Definition reduce_12 :=
+ Eval lazy beta iota delta[reduce_n1] in
+ reduce_n1 _ _ zero w11_eq0 reduce_11 N12.
+ Definition reduce_13 :=
+ Eval lazy beta iota delta[reduce_n1] in
+ reduce_n1 _ _ zero w12_eq0 reduce_12 N13.
+ Definition reduce_14 :=
+ Eval lazy beta iota delta[reduce_n1] in
+ reduce_n1 _ _ zero w13_eq0 reduce_13 (Nn 0).
+ Definition reduce_n n :=
+ Eval lazy beta iota delta[reduce_n] in
+ reduce_n _ _ zero reduce_14 Nn n.
+
+ (***************************************************************)
+ (* *)
+ (* Successor *)
+ (* *)
+ (***************************************************************)
+
Definition w0_succ_c := w0_op.(znz_succ_c).
Definition w1_succ_c := w1_op.(znz_succ_c).
Definition w2_succ_c := w2_op.(znz_succ_c).
@@ -219,466 +1729,130 @@ Module Make (W0:W0Type).
end
end.
- Definition extend1 :=
- Eval lazy beta zeta iota delta [extend]in extend 1.
- Definition extend2 :=
- Eval lazy beta zeta iota delta [extend]in extend 2.
- Definition extend3 :=
- Eval lazy beta zeta iota delta [extend]in extend 3.
- Definition extend4 :=
- Eval lazy beta zeta iota delta [extend]in extend 4.
- Definition extend5 :=
- Eval lazy beta zeta iota delta [extend]in extend 5.
- Definition extend6 :=
- Eval lazy beta zeta iota delta [extend]in extend 6.
- Definition extend7 :=
- Eval lazy beta zeta iota delta [extend]in extend 7.
- Definition extend8 :=
- Eval lazy beta zeta iota delta [extend]in extend 8.
- Definition extend9 :=
- Eval lazy beta zeta iota delta [extend]in extend 9.
- Definition extend10 :=
- Eval lazy beta zeta iota delta [extend]in extend 10.
- Definition extend11 :=
- Eval lazy beta zeta iota delta [extend]in extend 11.
- Definition extend12 :=
- Eval lazy beta zeta iota delta [extend]in extend 12.
- Definition extend13 :=
- Eval lazy beta zeta iota delta [extend]in extend 13.
-
- Definition w0_eq0 := w0_op.(znz_eq0).
- Definition w1_eq0 := w1_op.(znz_eq0).
- Definition w2_eq0 := w2_op.(znz_eq0).
- Definition w3_eq0 := w3_op.(znz_eq0).
- Definition w4_eq0 := w4_op.(znz_eq0).
- Definition w5_eq0 := w5_op.(znz_eq0).
- Definition w6_eq0 := w6_op.(znz_eq0).
- Definition w7_eq0 := w7_op.(znz_eq0).
- Definition w8_eq0 := w8_op.(znz_eq0).
- Definition w9_eq0 := w9_op.(znz_eq0).
- Definition w10_eq0 := w10_op.(znz_eq0).
- Definition w11_eq0 := w11_op.(znz_eq0).
- Definition w12_eq0 := w12_op.(znz_eq0).
- Definition w13_eq0 := w13_op.(znz_eq0).
-
- Definition w0_0W := w0_op.(znz_0W).
- Definition w1_0W := w1_op.(znz_0W).
- Definition w2_0W := w2_op.(znz_0W).
- Definition w3_0W := w3_op.(znz_0W).
- Definition w4_0W := w4_op.(znz_0W).
- Definition w5_0W := w5_op.(znz_0W).
- Definition w6_0W := w6_op.(znz_0W).
- Definition w7_0W := w7_op.(znz_0W).
- Definition w8_0W := w8_op.(znz_0W).
- Definition w9_0W := w9_op.(znz_0W).
- Definition w10_0W := w10_op.(znz_0W).
- Definition w11_0W := w11_op.(znz_0W).
- Definition w12_0W := w12_op.(znz_0W).
- Definition w13_0W := w13_op.(znz_0W).
+ Theorem succ_spec: forall n, [succ n] = [n] + 1.
+ Admitted.
- Definition w0_WW := w0_op.(znz_WW).
-
- Definition w0_add_c := w0_op.(znz_add_c).
- Definition w1_add_c := w1_op.(znz_add_c).
- Definition w2_add_c := w2_op.(znz_add_c).
- Definition w3_add_c := w3_op.(znz_add_c).
- Definition w4_add_c := w4_op.(znz_add_c).
- Definition w5_add_c := w5_op.(znz_add_c).
- Definition w6_add_c := w6_op.(znz_add_c).
- Definition w7_add_c := w7_op.(znz_add_c).
- Definition w8_add_c := w8_op.(znz_add_c).
- Definition w9_add_c := w9_op.(znz_add_c).
- Definition w10_add_c := w10_op.(znz_add_c).
- Definition w11_add_c := w11_op.(znz_add_c).
- Definition w12_add_c := w12_op.(znz_add_c).
- Definition w13_add_c := w13_op.(znz_add_c).
+ (***************************************************************)
+ (* *)
+ (* Adddition *)
+ (* *)
+ (***************************************************************)
+ Definition w0_add_c := znz_add_c w0_op.
Definition w0_add x y :=
match w0_add_c x y with
| C0 r => N0 r
| C1 r => N1 (WW one0 r)
end.
+
+ Definition w1_add_c := znz_add_c w1_op.
Definition w1_add x y :=
match w1_add_c x y with
| C0 r => N1 r
| C1 r => N2 (WW one1 r)
end.
+
+ Definition w2_add_c := znz_add_c w2_op.
Definition w2_add x y :=
match w2_add_c x y with
| C0 r => N2 r
| C1 r => N3 (WW one2 r)
end.
+
+ Definition w3_add_c := znz_add_c w3_op.
Definition w3_add x y :=
match w3_add_c x y with
| C0 r => N3 r
| C1 r => N4 (WW one3 r)
end.
+
+ Definition w4_add_c := znz_add_c w4_op.
Definition w4_add x y :=
match w4_add_c x y with
| C0 r => N4 r
| C1 r => N5 (WW one4 r)
end.
+
+ Definition w5_add_c := znz_add_c w5_op.
Definition w5_add x y :=
match w5_add_c x y with
| C0 r => N5 r
| C1 r => N6 (WW one5 r)
end.
+
+ Definition w6_add_c := znz_add_c w6_op.
Definition w6_add x y :=
match w6_add_c x y with
| C0 r => N6 r
| C1 r => N7 (WW one6 r)
end.
+
+ Definition w7_add_c := znz_add_c w7_op.
Definition w7_add x y :=
match w7_add_c x y with
| C0 r => N7 r
| C1 r => N8 (WW one7 r)
end.
+
+ Definition w8_add_c := znz_add_c w8_op.
Definition w8_add x y :=
match w8_add_c x y with
| C0 r => N8 r
| C1 r => N9 (WW one8 r)
end.
+
+ Definition w9_add_c := znz_add_c w9_op.
Definition w9_add x y :=
match w9_add_c x y with
| C0 r => N9 r
| C1 r => N10 (WW one9 r)
end.
+
+ Definition w10_add_c := znz_add_c w10_op.
Definition w10_add x y :=
match w10_add_c x y with
| C0 r => N10 r
| C1 r => N11 (WW one10 r)
end.
+
+ Definition w11_add_c := znz_add_c w11_op.
Definition w11_add x y :=
match w11_add_c x y with
| C0 r => N11 r
| C1 r => N12 (WW one11 r)
end.
+
+ Definition w12_add_c := znz_add_c w12_op.
Definition w12_add x y :=
match w12_add_c x y with
| C0 r => N12 r
| C1 r => N13 (WW one12 r)
end.
+
+ Definition w13_add_c := znz_add_c w13_op.
Definition w13_add x y :=
match w13_add_c x y with
| C0 r => N13 r
| C1 r => Nn 0 (WW one13 r)
end.
+
Definition addn n (x y : word w13 (S n)) :=
let op := make_op n in
match op.(znz_add_c) x y with
| C0 r => Nn n r
| C1 r => Nn (S n) (WW op.(znz_1) r) end.
- Definition add x y :=
- match x, y with
- | N0 wx, N0 wy => w0_add wx wy
- | N0 wx, N1 wy =>
- if w0_eq0 wx then y else w1_add (WW w_0 wx) wy
- | N0 wx, N2 wy =>
- if w0_eq0 wx then y else w2_add (extend1 w0 (WW w_0 wx)) wy
- | N0 wx, N3 wy =>
- if w0_eq0 wx then y else w3_add (extend2 w0 (WW w_0 wx)) wy
- | N0 wx, N4 wy =>
- if w0_eq0 wx then y else w4_add (extend3 w0 (WW w_0 wx)) wy
- | N0 wx, N5 wy =>
- if w0_eq0 wx then y else w5_add (extend4 w0 (WW w_0 wx)) wy
- | N0 wx, N6 wy =>
- if w0_eq0 wx then y else w6_add (extend5 w0 (WW w_0 wx)) wy
- | N0 wx, N7 wy =>
- if w0_eq0 wx then y else w7_add (extend6 w0 (WW w_0 wx)) wy
- | N0 wx, N8 wy =>
- if w0_eq0 wx then y else w8_add (extend7 w0 (WW w_0 wx)) wy
- | N0 wx, N9 wy =>
- if w0_eq0 wx then y else w9_add (extend8 w0 (WW w_0 wx)) wy
- | N0 wx, N10 wy =>
- if w0_eq0 wx then y else w10_add (extend9 w0 (WW w_0 wx)) wy
- | N0 wx, N11 wy =>
- if w0_eq0 wx then y else w11_add (extend10 w0 (WW w_0 wx)) wy
- | N0 wx, N12 wy =>
- if w0_eq0 wx then y else w12_add (extend11 w0 (WW w_0 wx)) wy
- | N0 wx, N13 wy =>
- if w0_eq0 wx then y else w13_add (extend12 w0 (WW w_0 wx)) wy
- | N0 wx, Nn n wy =>
- if w0_eq0 wx then y
- else addn n (extend n w13 (extend13 w0 (WW w_0 wx))) wy
- | N1 wx, N0 wy =>
- if w0_eq0 wy then x else w1_add wx (WW w_0 wy)
- | N1 wx, N1 wy => w1_add wx wy
- | N1 wx, N2 wy => w2_add (extend1 w0 wx) wy
- | N1 wx, N3 wy => w3_add (extend2 w0 wx) wy
- | N1 wx, N4 wy => w4_add (extend3 w0 wx) wy
- | N1 wx, N5 wy => w5_add (extend4 w0 wx) wy
- | N1 wx, N6 wy => w6_add (extend5 w0 wx) wy
- | N1 wx, N7 wy => w7_add (extend6 w0 wx) wy
- | N1 wx, N8 wy => w8_add (extend7 w0 wx) wy
- | N1 wx, N9 wy => w9_add (extend8 w0 wx) wy
- | N1 wx, N10 wy => w10_add (extend9 w0 wx) wy
- | N1 wx, N11 wy => w11_add (extend10 w0 wx) wy
- | N1 wx, N12 wy => w12_add (extend11 w0 wx) wy
- | N1 wx, N13 wy => w13_add (extend12 w0 wx) wy
- | N1 wx, Nn n wy => addn n (extend n w13 (extend13 w0 wx)) wy
- | N2 wx, N0 wy =>
- if w0_eq0 wy then x else w2_add wx (extend1 w0 (WW w_0 wy))
- | N2 wx, N1 wy => w2_add wx (extend1 w0 wy)
- | N2 wx, N2 wy => w2_add wx wy
- | N2 wx, N3 wy => w3_add (extend1 w1 wx) wy
- | N2 wx, N4 wy => w4_add (extend2 w1 wx) wy
- | N2 wx, N5 wy => w5_add (extend3 w1 wx) wy
- | N2 wx, N6 wy => w6_add (extend4 w1 wx) wy
- | N2 wx, N7 wy => w7_add (extend5 w1 wx) wy
- | N2 wx, N8 wy => w8_add (extend6 w1 wx) wy
- | N2 wx, N9 wy => w9_add (extend7 w1 wx) wy
- | N2 wx, N10 wy => w10_add (extend8 w1 wx) wy
- | N2 wx, N11 wy => w11_add (extend9 w1 wx) wy
- | N2 wx, N12 wy => w12_add (extend10 w1 wx) wy
- | N2 wx, N13 wy => w13_add (extend11 w1 wx) wy
- | N2 wx, Nn n wy => addn n (extend n w13 (extend12 w1 wx)) wy
- | N3 wx, N0 wy =>
- if w0_eq0 wy then x else w3_add wx (extend2 w0 (WW w_0 wy))
- | N3 wx, N1 wy => w3_add wx (extend2 w0 wy)
- | N3 wx, N2 wy => w3_add wx (extend1 w1 wy)
- | N3 wx, N3 wy => w3_add wx wy
- | N3 wx, N4 wy => w4_add (extend1 w2 wx) wy
- | N3 wx, N5 wy => w5_add (extend2 w2 wx) wy
- | N3 wx, N6 wy => w6_add (extend3 w2 wx) wy
- | N3 wx, N7 wy => w7_add (extend4 w2 wx) wy
- | N3 wx, N8 wy => w8_add (extend5 w2 wx) wy
- | N3 wx, N9 wy => w9_add (extend6 w2 wx) wy
- | N3 wx, N10 wy => w10_add (extend7 w2 wx) wy
- | N3 wx, N11 wy => w11_add (extend8 w2 wx) wy
- | N3 wx, N12 wy => w12_add (extend9 w2 wx) wy
- | N3 wx, N13 wy => w13_add (extend10 w2 wx) wy
- | N3 wx, Nn n wy => addn n (extend n w13 (extend11 w2 wx)) wy
- | N4 wx, N0 wy =>
- if w0_eq0 wy then x else w4_add wx (extend3 w0 (WW w_0 wy))
- | N4 wx, N1 wy => w4_add wx (extend3 w0 wy)
- | N4 wx, N2 wy => w4_add wx (extend2 w1 wy)
- | N4 wx, N3 wy => w4_add wx (extend1 w2 wy)
- | N4 wx, N4 wy => w4_add wx wy
- | N4 wx, N5 wy => w5_add (extend1 w3 wx) wy
- | N4 wx, N6 wy => w6_add (extend2 w3 wx) wy
- | N4 wx, N7 wy => w7_add (extend3 w3 wx) wy
- | N4 wx, N8 wy => w8_add (extend4 w3 wx) wy
- | N4 wx, N9 wy => w9_add (extend5 w3 wx) wy
- | N4 wx, N10 wy => w10_add (extend6 w3 wx) wy
- | N4 wx, N11 wy => w11_add (extend7 w3 wx) wy
- | N4 wx, N12 wy => w12_add (extend8 w3 wx) wy
- | N4 wx, N13 wy => w13_add (extend9 w3 wx) wy
- | N4 wx, Nn n wy => addn n (extend n w13 (extend10 w3 wx)) wy
- | N5 wx, N0 wy =>
- if w0_eq0 wy then x else w5_add wx (extend4 w0 (WW w_0 wy))
- | N5 wx, N1 wy => w5_add wx (extend4 w0 wy)
- | N5 wx, N2 wy => w5_add wx (extend3 w1 wy)
- | N5 wx, N3 wy => w5_add wx (extend2 w2 wy)
- | N5 wx, N4 wy => w5_add wx (extend1 w3 wy)
- | N5 wx, N5 wy => w5_add wx wy
- | N5 wx, N6 wy => w6_add (extend1 w4 wx) wy
- | N5 wx, N7 wy => w7_add (extend2 w4 wx) wy
- | N5 wx, N8 wy => w8_add (extend3 w4 wx) wy
- | N5 wx, N9 wy => w9_add (extend4 w4 wx) wy
- | N5 wx, N10 wy => w10_add (extend5 w4 wx) wy
- | N5 wx, N11 wy => w11_add (extend6 w4 wx) wy
- | N5 wx, N12 wy => w12_add (extend7 w4 wx) wy
- | N5 wx, N13 wy => w13_add (extend8 w4 wx) wy
- | N5 wx, Nn n wy => addn n (extend n w13 (extend9 w4 wx)) wy
- | N6 wx, N0 wy =>
- if w0_eq0 wy then x else w6_add wx (extend5 w0 (WW w_0 wy))
- | N6 wx, N1 wy => w6_add wx (extend5 w0 wy)
- | N6 wx, N2 wy => w6_add wx (extend4 w1 wy)
- | N6 wx, N3 wy => w6_add wx (extend3 w2 wy)
- | N6 wx, N4 wy => w6_add wx (extend2 w3 wy)
- | N6 wx, N5 wy => w6_add wx (extend1 w4 wy)
- | N6 wx, N6 wy => w6_add wx wy
- | N6 wx, N7 wy => w7_add (extend1 w5 wx) wy
- | N6 wx, N8 wy => w8_add (extend2 w5 wx) wy
- | N6 wx, N9 wy => w9_add (extend3 w5 wx) wy
- | N6 wx, N10 wy => w10_add (extend4 w5 wx) wy
- | N6 wx, N11 wy => w11_add (extend5 w5 wx) wy
- | N6 wx, N12 wy => w12_add (extend6 w5 wx) wy
- | N6 wx, N13 wy => w13_add (extend7 w5 wx) wy
- | N6 wx, Nn n wy => addn n (extend n w13 (extend8 w5 wx)) wy
- | N7 wx, N0 wy =>
- if w0_eq0 wy then x else w7_add wx (extend6 w0 (WW w_0 wy))
- | N7 wx, N1 wy => w7_add wx (extend6 w0 wy)
- | N7 wx, N2 wy => w7_add wx (extend5 w1 wy)
- | N7 wx, N3 wy => w7_add wx (extend4 w2 wy)
- | N7 wx, N4 wy => w7_add wx (extend3 w3 wy)
- | N7 wx, N5 wy => w7_add wx (extend2 w4 wy)
- | N7 wx, N6 wy => w7_add wx (extend1 w5 wy)
- | N7 wx, N7 wy => w7_add wx wy
- | N7 wx, N8 wy => w8_add (extend1 w6 wx) wy
- | N7 wx, N9 wy => w9_add (extend2 w6 wx) wy
- | N7 wx, N10 wy => w10_add (extend3 w6 wx) wy
- | N7 wx, N11 wy => w11_add (extend4 w6 wx) wy
- | N7 wx, N12 wy => w12_add (extend5 w6 wx) wy
- | N7 wx, N13 wy => w13_add (extend6 w6 wx) wy
- | N7 wx, Nn n wy => addn n (extend n w13 (extend7 w6 wx)) wy
- | N8 wx, N0 wy =>
- if w0_eq0 wy then x else w8_add wx (extend7 w0 (WW w_0 wy))
- | N8 wx, N1 wy => w8_add wx (extend7 w0 wy)
- | N8 wx, N2 wy => w8_add wx (extend6 w1 wy)
- | N8 wx, N3 wy => w8_add wx (extend5 w2 wy)
- | N8 wx, N4 wy => w8_add wx (extend4 w3 wy)
- | N8 wx, N5 wy => w8_add wx (extend3 w4 wy)
- | N8 wx, N6 wy => w8_add wx (extend2 w5 wy)
- | N8 wx, N7 wy => w8_add wx (extend1 w6 wy)
- | N8 wx, N8 wy => w8_add wx wy
- | N8 wx, N9 wy => w9_add (extend1 w7 wx) wy
- | N8 wx, N10 wy => w10_add (extend2 w7 wx) wy
- | N8 wx, N11 wy => w11_add (extend3 w7 wx) wy
- | N8 wx, N12 wy => w12_add (extend4 w7 wx) wy
- | N8 wx, N13 wy => w13_add (extend5 w7 wx) wy
- | N8 wx, Nn n wy => addn n (extend n w13 (extend6 w7 wx)) wy
- | N9 wx, N0 wy =>
- if w0_eq0 wy then x else w9_add wx (extend8 w0 (WW w_0 wy))
- | N9 wx, N1 wy => w9_add wx (extend8 w0 wy)
- | N9 wx, N2 wy => w9_add wx (extend7 w1 wy)
- | N9 wx, N3 wy => w9_add wx (extend6 w2 wy)
- | N9 wx, N4 wy => w9_add wx (extend5 w3 wy)
- | N9 wx, N5 wy => w9_add wx (extend4 w4 wy)
- | N9 wx, N6 wy => w9_add wx (extend3 w5 wy)
- | N9 wx, N7 wy => w9_add wx (extend2 w6 wy)
- | N9 wx, N8 wy => w9_add wx (extend1 w7 wy)
- | N9 wx, N9 wy => w9_add wx wy
- | N9 wx, N10 wy => w10_add (extend1 w8 wx) wy
- | N9 wx, N11 wy => w11_add (extend2 w8 wx) wy
- | N9 wx, N12 wy => w12_add (extend3 w8 wx) wy
- | N9 wx, N13 wy => w13_add (extend4 w8 wx) wy
- | N9 wx, Nn n wy => addn n (extend n w13 (extend5 w8 wx)) wy
- | N10 wx, N0 wy =>
- if w0_eq0 wy then x else w10_add wx (extend9 w0 (WW w_0 wy))
- | N10 wx, N1 wy => w10_add wx (extend9 w0 wy)
- | N10 wx, N2 wy => w10_add wx (extend8 w1 wy)
- | N10 wx, N3 wy => w10_add wx (extend7 w2 wy)
- | N10 wx, N4 wy => w10_add wx (extend6 w3 wy)
- | N10 wx, N5 wy => w10_add wx (extend5 w4 wy)
- | N10 wx, N6 wy => w10_add wx (extend4 w5 wy)
- | N10 wx, N7 wy => w10_add wx (extend3 w6 wy)
- | N10 wx, N8 wy => w10_add wx (extend2 w7 wy)
- | N10 wx, N9 wy => w10_add wx (extend1 w8 wy)
- | N10 wx, N10 wy => w10_add wx wy
- | N10 wx, N11 wy => w11_add (extend1 w9 wx) wy
- | N10 wx, N12 wy => w12_add (extend2 w9 wx) wy
- | N10 wx, N13 wy => w13_add (extend3 w9 wx) wy
- | N10 wx, Nn n wy => addn n (extend n w13 (extend4 w9 wx)) wy
- | N11 wx, N0 wy =>
- if w0_eq0 wy then x else w11_add wx (extend10 w0 (WW w_0 wy))
- | N11 wx, N1 wy => w11_add wx (extend10 w0 wy)
- | N11 wx, N2 wy => w11_add wx (extend9 w1 wy)
- | N11 wx, N3 wy => w11_add wx (extend8 w2 wy)
- | N11 wx, N4 wy => w11_add wx (extend7 w3 wy)
- | N11 wx, N5 wy => w11_add wx (extend6 w4 wy)
- | N11 wx, N6 wy => w11_add wx (extend5 w5 wy)
- | N11 wx, N7 wy => w11_add wx (extend4 w6 wy)
- | N11 wx, N8 wy => w11_add wx (extend3 w7 wy)
- | N11 wx, N9 wy => w11_add wx (extend2 w8 wy)
- | N11 wx, N10 wy => w11_add wx (extend1 w9 wy)
- | N11 wx, N11 wy => w11_add wx wy
- | N11 wx, N12 wy => w12_add (extend1 w10 wx) wy
- | N11 wx, N13 wy => w13_add (extend2 w10 wx) wy
- | N11 wx, Nn n wy => addn n (extend n w13 (extend3 w10 wx)) wy
- | N12 wx, N0 wy =>
- if w0_eq0 wy then x else w12_add wx (extend11 w0 (WW w_0 wy))
- | N12 wx, N1 wy => w12_add wx (extend11 w0 wy)
- | N12 wx, N2 wy => w12_add wx (extend10 w1 wy)
- | N12 wx, N3 wy => w12_add wx (extend9 w2 wy)
- | N12 wx, N4 wy => w12_add wx (extend8 w3 wy)
- | N12 wx, N5 wy => w12_add wx (extend7 w4 wy)
- | N12 wx, N6 wy => w12_add wx (extend6 w5 wy)
- | N12 wx, N7 wy => w12_add wx (extend5 w6 wy)
- | N12 wx, N8 wy => w12_add wx (extend4 w7 wy)
- | N12 wx, N9 wy => w12_add wx (extend3 w8 wy)
- | N12 wx, N10 wy => w12_add wx (extend2 w9 wy)
- | N12 wx, N11 wy => w12_add wx (extend1 w10 wy)
- | N12 wx, N12 wy => w12_add wx wy
- | N12 wx, N13 wy => w13_add (extend1 w11 wx) wy
- | N12 wx, Nn n wy => addn n (extend n w13 (extend2 w11 wx)) wy
- | N13 wx, N0 wy =>
- if w0_eq0 wy then x else w13_add wx (extend12 w0 (WW w_0 wy))
- | N13 wx, N1 wy => w13_add wx (extend12 w0 wy)
- | N13 wx, N2 wy => w13_add wx (extend11 w1 wy)
- | N13 wx, N3 wy => w13_add wx (extend10 w2 wy)
- | N13 wx, N4 wy => w13_add wx (extend9 w3 wy)
- | N13 wx, N5 wy => w13_add wx (extend8 w4 wy)
- | N13 wx, N6 wy => w13_add wx (extend7 w5 wy)
- | N13 wx, N7 wy => w13_add wx (extend6 w6 wy)
- | N13 wx, N8 wy => w13_add wx (extend5 w7 wy)
- | N13 wx, N9 wy => w13_add wx (extend4 w8 wy)
- | N13 wx, N10 wy => w13_add wx (extend3 w9 wy)
- | N13 wx, N11 wy => w13_add wx (extend2 w10 wy)
- | N13 wx, N12 wy => w13_add wx (extend1 w11 wy)
- | N13 wx, N13 wy => w13_add wx wy
- | N13 wx, Nn n wy => addn n (extend n w13 (extend1 w12 wx)) wy
- | Nn n wx, N0 wy =>
- if w0_eq0 wy then x
- else addn n wx (extend n w13 (extend13 w0 (WW w_0 wy)))
- | Nn n wx, N1 wy => addn n wx (extend n w13 (extend13 w0 wy))
- | Nn n wx, N2 wy => addn n wx (extend n w13 (extend12 w1 wy))
- | Nn n wx, N3 wy => addn n wx (extend n w13 (extend11 w2 wy))
- | Nn n wx, N4 wy => addn n wx (extend n w13 (extend10 w3 wy))
- | Nn n wx, N5 wy => addn n wx (extend n w13 (extend9 w4 wy))
- | Nn n wx, N6 wy => addn n wx (extend n w13 (extend8 w5 wy))
- | Nn n wx, N7 wy => addn n wx (extend n w13 (extend7 w6 wy))
- | Nn n wx, N8 wy => addn n wx (extend n w13 (extend6 w7 wy))
- | Nn n wx, N9 wy => addn n wx (extend n w13 (extend5 w8 wy))
- | Nn n wx, N10 wy => addn n wx (extend n w13 (extend4 w9 wy))
- | Nn n wx, N11 wy => addn n wx (extend n w13 (extend3 w10 wy))
- | Nn n wx, N12 wy => addn n wx (extend n w13 (extend2 w11 wy))
- | Nn n wx, N13 wy => addn n wx (extend n w13 (extend1 w12 wy))
- | Nn n wx, Nn m wy =>
- let mn := Max.max n m in
- let d := diff n m in
- addn mn
- (castm (diff_r n m) (extend_tr wx (snd d)))
- (castm (diff_l n m) (extend_tr wy (fst d)))
- end.
+ Definition add := Eval lazy beta delta [same_level] in
+ (same_level t_ w0_add w1_add w2_add w3_add w4_add w5_add w6_add w7_add w8_add w9_add w10_add w11_add w12_add w13_add addn).
- Definition reduce_0 (x:w) := N0 x.
- Definition reduce_1 :=
- Eval lazy beta iota delta[reduce_n1] in
- reduce_n1 _ _ zero w0_eq0 N0 N1.
- Definition reduce_2 :=
- Eval lazy beta iota delta[reduce_n1] in
- reduce_n1 _ _ zero w1_eq0 reduce_1 N2.
- Definition reduce_3 :=
- Eval lazy beta iota delta[reduce_n1] in
- reduce_n1 _ _ zero w2_eq0 reduce_2 N3.
- Definition reduce_4 :=
- Eval lazy beta iota delta[reduce_n1] in
- reduce_n1 _ _ zero w3_eq0 reduce_3 N4.
- Definition reduce_5 :=
- Eval lazy beta iota delta[reduce_n1] in
- reduce_n1 _ _ zero w4_eq0 reduce_4 N5.
- Definition reduce_6 :=
- Eval lazy beta iota delta[reduce_n1] in
- reduce_n1 _ _ zero w5_eq0 reduce_5 N6.
- Definition reduce_7 :=
- Eval lazy beta iota delta[reduce_n1] in
- reduce_n1 _ _ zero w6_eq0 reduce_6 N7.
- Definition reduce_8 :=
- Eval lazy beta iota delta[reduce_n1] in
- reduce_n1 _ _ zero w7_eq0 reduce_7 N8.
- Definition reduce_9 :=
- Eval lazy beta iota delta[reduce_n1] in
- reduce_n1 _ _ zero w8_eq0 reduce_8 N9.
- Definition reduce_10 :=
- Eval lazy beta iota delta[reduce_n1] in
- reduce_n1 _ _ zero w9_eq0 reduce_9 N10.
- Definition reduce_11 :=
- Eval lazy beta iota delta[reduce_n1] in
- reduce_n1 _ _ zero w10_eq0 reduce_10 N11.
- Definition reduce_12 :=
- Eval lazy beta iota delta[reduce_n1] in
- reduce_n1 _ _ zero w11_eq0 reduce_11 N12.
- Definition reduce_13 :=
- Eval lazy beta iota delta[reduce_n1] in
- reduce_n1 _ _ zero w12_eq0 reduce_12 N13.
- Definition reduce_14 :=
- Eval lazy beta iota delta[reduce_n1] in
- reduce_n1 _ _ zero w13_eq0 reduce_13 (Nn 0).
- Definition reduce_n n :=
- Eval lazy beta iota delta[reduce_n] in
- reduce_n _ _ zero reduce_14 Nn n.
+ Theorem spec_add: forall x y, [add x y] = [x] + [y].
+ Admitted.
+
+ (***************************************************************)
+ (* *)
+ (* Predecessor *)
+ (* *)
+ (***************************************************************)
Definition w0_pred_c := w0_op.(znz_pred_c).
Definition w1_pred_c := w1_op.(znz_pred_c).
@@ -775,6 +1949,14 @@ Module Make (W0:W0Type).
end
end.
+ Let spec_pred: forall x, 0 < [x] -> [pred x] = [x] - 1.
+ Admitted.
+
+ (***************************************************************)
+ (* *)
+ (* Subtraction *)
+ (* *)
+ (***************************************************************)
Definition w0_sub_c := w0_op.(znz_sub_c).
Definition w1_sub_c := w1_op.(znz_sub_c).
@@ -868,282 +2050,20 @@ Module Make (W0:W0Type).
| C0 r => Nn n r
| C1 r => N0 w_0 end.
- Definition sub x y :=
- match x, y with
- | N0 wx, N0 wy => w0_sub wx wy
- | N0 wx, N1 wy =>
- if w0_eq0 wx then zero else w1_sub (WW w_0 wx) wy
- | N0 wx, N2 wy =>
- if w0_eq0 wx then zero else w2_sub (extend1 w0 (WW w_0 wx)) wy
- | N0 wx, N3 wy =>
- if w0_eq0 wx then zero else w3_sub (extend2 w0 (WW w_0 wx)) wy
- | N0 wx, N4 wy =>
- if w0_eq0 wx then zero else w4_sub (extend3 w0 (WW w_0 wx)) wy
- | N0 wx, N5 wy =>
- if w0_eq0 wx then zero else w5_sub (extend4 w0 (WW w_0 wx)) wy
- | N0 wx, N6 wy =>
- if w0_eq0 wx then zero else w6_sub (extend5 w0 (WW w_0 wx)) wy
- | N0 wx, N7 wy =>
- if w0_eq0 wx then zero else w7_sub (extend6 w0 (WW w_0 wx)) wy
- | N0 wx, N8 wy =>
- if w0_eq0 wx then zero else w8_sub (extend7 w0 (WW w_0 wx)) wy
- | N0 wx, N9 wy =>
- if w0_eq0 wx then zero else w9_sub (extend8 w0 (WW w_0 wx)) wy
- | N0 wx, N10 wy =>
- if w0_eq0 wx then zero else w10_sub (extend9 w0 (WW w_0 wx)) wy
- | N0 wx, N11 wy =>
- if w0_eq0 wx then zero else w11_sub (extend10 w0 (WW w_0 wx)) wy
- | N0 wx, N12 wy =>
- if w0_eq0 wx then zero else w12_sub (extend11 w0 (WW w_0 wx)) wy
- | N0 wx, N13 wy =>
- if w0_eq0 wx then zero else w13_sub (extend12 w0 (WW w_0 wx)) wy
- | N0 wx, Nn n wy =>
- if w0_eq0 wx then zero
- else subn n (extend n w13 (extend13 w0 (WW w_0 wx))) wy
- | N1 wx, N0 wy =>
- if w0_eq0 wy then x
- else w1_sub wx (WW w_0 wy)
- | N1 wx, N1 wy => w1_sub wx wy
- | N1 wx, N2 wy => w2_sub (extend1 w0 wx) wy
- | N1 wx, N3 wy => w3_sub (extend2 w0 wx) wy
- | N1 wx, N4 wy => w4_sub (extend3 w0 wx) wy
- | N1 wx, N5 wy => w5_sub (extend4 w0 wx) wy
- | N1 wx, N6 wy => w6_sub (extend5 w0 wx) wy
- | N1 wx, N7 wy => w7_sub (extend6 w0 wx) wy
- | N1 wx, N8 wy => w8_sub (extend7 w0 wx) wy
- | N1 wx, N9 wy => w9_sub (extend8 w0 wx) wy
- | N1 wx, N10 wy => w10_sub (extend9 w0 wx) wy
- | N1 wx, N11 wy => w11_sub (extend10 w0 wx) wy
- | N1 wx, N12 wy => w12_sub (extend11 w0 wx) wy
- | N1 wx, N13 wy => w13_sub (extend12 w0 wx) wy
- | N1 wx, Nn n wy => subn n (extend n w13 (extend13 w0 wx)) wy
- | N2 wx, N0 wy =>
- if w0_eq0 wy then x
- else w2_sub wx (extend1 w0 (WW w_0 wy))
- | N2 wx, N1 wy => w2_sub wx (extend1 w0 wy)
- | N2 wx, N2 wy => w2_sub wx wy
- | N2 wx, N3 wy => w3_sub (extend1 w1 wx) wy
- | N2 wx, N4 wy => w4_sub (extend2 w1 wx) wy
- | N2 wx, N5 wy => w5_sub (extend3 w1 wx) wy
- | N2 wx, N6 wy => w6_sub (extend4 w1 wx) wy
- | N2 wx, N7 wy => w7_sub (extend5 w1 wx) wy
- | N2 wx, N8 wy => w8_sub (extend6 w1 wx) wy
- | N2 wx, N9 wy => w9_sub (extend7 w1 wx) wy
- | N2 wx, N10 wy => w10_sub (extend8 w1 wx) wy
- | N2 wx, N11 wy => w11_sub (extend9 w1 wx) wy
- | N2 wx, N12 wy => w12_sub (extend10 w1 wx) wy
- | N2 wx, N13 wy => w13_sub (extend11 w1 wx) wy
- | N2 wx, Nn n wy => subn n (extend n w13 (extend12 w1 wx)) wy
- | N3 wx, N0 wy =>
- if w0_eq0 wy then x
- else w3_sub wx (extend2 w0 (WW w_0 wy))
- | N3 wx, N1 wy => w3_sub wx (extend2 w0 wy)
- | N3 wx, N2 wy => w3_sub wx (extend1 w1 wy)
- | N3 wx, N3 wy => w3_sub wx wy
- | N3 wx, N4 wy => w4_sub (extend1 w2 wx) wy
- | N3 wx, N5 wy => w5_sub (extend2 w2 wx) wy
- | N3 wx, N6 wy => w6_sub (extend3 w2 wx) wy
- | N3 wx, N7 wy => w7_sub (extend4 w2 wx) wy
- | N3 wx, N8 wy => w8_sub (extend5 w2 wx) wy
- | N3 wx, N9 wy => w9_sub (extend6 w2 wx) wy
- | N3 wx, N10 wy => w10_sub (extend7 w2 wx) wy
- | N3 wx, N11 wy => w11_sub (extend8 w2 wx) wy
- | N3 wx, N12 wy => w12_sub (extend9 w2 wx) wy
- | N3 wx, N13 wy => w13_sub (extend10 w2 wx) wy
- | N3 wx, Nn n wy => subn n (extend n w13 (extend11 w2 wx)) wy
- | N4 wx, N0 wy =>
- if w0_eq0 wy then x
- else w4_sub wx (extend3 w0 (WW w_0 wy))
- | N4 wx, N1 wy => w4_sub wx (extend3 w0 wy)
- | N4 wx, N2 wy => w4_sub wx (extend2 w1 wy)
- | N4 wx, N3 wy => w4_sub wx (extend1 w2 wy)
- | N4 wx, N4 wy => w4_sub wx wy
- | N4 wx, N5 wy => w5_sub (extend1 w3 wx) wy
- | N4 wx, N6 wy => w6_sub (extend2 w3 wx) wy
- | N4 wx, N7 wy => w7_sub (extend3 w3 wx) wy
- | N4 wx, N8 wy => w8_sub (extend4 w3 wx) wy
- | N4 wx, N9 wy => w9_sub (extend5 w3 wx) wy
- | N4 wx, N10 wy => w10_sub (extend6 w3 wx) wy
- | N4 wx, N11 wy => w11_sub (extend7 w3 wx) wy
- | N4 wx, N12 wy => w12_sub (extend8 w3 wx) wy
- | N4 wx, N13 wy => w13_sub (extend9 w3 wx) wy
- | N4 wx, Nn n wy => subn n (extend n w13 (extend10 w3 wx)) wy
- | N5 wx, N0 wy =>
- if w0_eq0 wy then x
- else w5_sub wx (extend4 w0 (WW w_0 wy))
- | N5 wx, N1 wy => w5_sub wx (extend4 w0 wy)
- | N5 wx, N2 wy => w5_sub wx (extend3 w1 wy)
- | N5 wx, N3 wy => w5_sub wx (extend2 w2 wy)
- | N5 wx, N4 wy => w5_sub wx (extend1 w3 wy)
- | N5 wx, N5 wy => w5_sub wx wy
- | N5 wx, N6 wy => w6_sub (extend1 w4 wx) wy
- | N5 wx, N7 wy => w7_sub (extend2 w4 wx) wy
- | N5 wx, N8 wy => w8_sub (extend3 w4 wx) wy
- | N5 wx, N9 wy => w9_sub (extend4 w4 wx) wy
- | N5 wx, N10 wy => w10_sub (extend5 w4 wx) wy
- | N5 wx, N11 wy => w11_sub (extend6 w4 wx) wy
- | N5 wx, N12 wy => w12_sub (extend7 w4 wx) wy
- | N5 wx, N13 wy => w13_sub (extend8 w4 wx) wy
- | N5 wx, Nn n wy => subn n (extend n w13 (extend9 w4 wx)) wy
- | N6 wx, N0 wy =>
- if w0_eq0 wy then x
- else w6_sub wx (extend5 w0 (WW w_0 wy))
- | N6 wx, N1 wy => w6_sub wx (extend5 w0 wy)
- | N6 wx, N2 wy => w6_sub wx (extend4 w1 wy)
- | N6 wx, N3 wy => w6_sub wx (extend3 w2 wy)
- | N6 wx, N4 wy => w6_sub wx (extend2 w3 wy)
- | N6 wx, N5 wy => w6_sub wx (extend1 w4 wy)
- | N6 wx, N6 wy => w6_sub wx wy
- | N6 wx, N7 wy => w7_sub (extend1 w5 wx) wy
- | N6 wx, N8 wy => w8_sub (extend2 w5 wx) wy
- | N6 wx, N9 wy => w9_sub (extend3 w5 wx) wy
- | N6 wx, N10 wy => w10_sub (extend4 w5 wx) wy
- | N6 wx, N11 wy => w11_sub (extend5 w5 wx) wy
- | N6 wx, N12 wy => w12_sub (extend6 w5 wx) wy
- | N6 wx, N13 wy => w13_sub (extend7 w5 wx) wy
- | N6 wx, Nn n wy => subn n (extend n w13 (extend8 w5 wx)) wy
- | N7 wx, N0 wy =>
- if w0_eq0 wy then x
- else w7_sub wx (extend6 w0 (WW w_0 wy))
- | N7 wx, N1 wy => w7_sub wx (extend6 w0 wy)
- | N7 wx, N2 wy => w7_sub wx (extend5 w1 wy)
- | N7 wx, N3 wy => w7_sub wx (extend4 w2 wy)
- | N7 wx, N4 wy => w7_sub wx (extend3 w3 wy)
- | N7 wx, N5 wy => w7_sub wx (extend2 w4 wy)
- | N7 wx, N6 wy => w7_sub wx (extend1 w5 wy)
- | N7 wx, N7 wy => w7_sub wx wy
- | N7 wx, N8 wy => w8_sub (extend1 w6 wx) wy
- | N7 wx, N9 wy => w9_sub (extend2 w6 wx) wy
- | N7 wx, N10 wy => w10_sub (extend3 w6 wx) wy
- | N7 wx, N11 wy => w11_sub (extend4 w6 wx) wy
- | N7 wx, N12 wy => w12_sub (extend5 w6 wx) wy
- | N7 wx, N13 wy => w13_sub (extend6 w6 wx) wy
- | N7 wx, Nn n wy => subn n (extend n w13 (extend7 w6 wx)) wy
- | N8 wx, N0 wy =>
- if w0_eq0 wy then x
- else w8_sub wx (extend7 w0 (WW w_0 wy))
- | N8 wx, N1 wy => w8_sub wx (extend7 w0 wy)
- | N8 wx, N2 wy => w8_sub wx (extend6 w1 wy)
- | N8 wx, N3 wy => w8_sub wx (extend5 w2 wy)
- | N8 wx, N4 wy => w8_sub wx (extend4 w3 wy)
- | N8 wx, N5 wy => w8_sub wx (extend3 w4 wy)
- | N8 wx, N6 wy => w8_sub wx (extend2 w5 wy)
- | N8 wx, N7 wy => w8_sub wx (extend1 w6 wy)
- | N8 wx, N8 wy => w8_sub wx wy
- | N8 wx, N9 wy => w9_sub (extend1 w7 wx) wy
- | N8 wx, N10 wy => w10_sub (extend2 w7 wx) wy
- | N8 wx, N11 wy => w11_sub (extend3 w7 wx) wy
- | N8 wx, N12 wy => w12_sub (extend4 w7 wx) wy
- | N8 wx, N13 wy => w13_sub (extend5 w7 wx) wy
- | N8 wx, Nn n wy => subn n (extend n w13 (extend6 w7 wx)) wy
- | N9 wx, N0 wy =>
- if w0_eq0 wy then x
- else w9_sub wx (extend8 w0 (WW w_0 wy))
- | N9 wx, N1 wy => w9_sub wx (extend8 w0 wy)
- | N9 wx, N2 wy => w9_sub wx (extend7 w1 wy)
- | N9 wx, N3 wy => w9_sub wx (extend6 w2 wy)
- | N9 wx, N4 wy => w9_sub wx (extend5 w3 wy)
- | N9 wx, N5 wy => w9_sub wx (extend4 w4 wy)
- | N9 wx, N6 wy => w9_sub wx (extend3 w5 wy)
- | N9 wx, N7 wy => w9_sub wx (extend2 w6 wy)
- | N9 wx, N8 wy => w9_sub wx (extend1 w7 wy)
- | N9 wx, N9 wy => w9_sub wx wy
- | N9 wx, N10 wy => w10_sub (extend1 w8 wx) wy
- | N9 wx, N11 wy => w11_sub (extend2 w8 wx) wy
- | N9 wx, N12 wy => w12_sub (extend3 w8 wx) wy
- | N9 wx, N13 wy => w13_sub (extend4 w8 wx) wy
- | N9 wx, Nn n wy => subn n (extend n w13 (extend5 w8 wx)) wy
- | N10 wx, N0 wy =>
- if w0_eq0 wy then x
- else w10_sub wx (extend9 w0 (WW w_0 wy))
- | N10 wx, N1 wy => w10_sub wx (extend9 w0 wy)
- | N10 wx, N2 wy => w10_sub wx (extend8 w1 wy)
- | N10 wx, N3 wy => w10_sub wx (extend7 w2 wy)
- | N10 wx, N4 wy => w10_sub wx (extend6 w3 wy)
- | N10 wx, N5 wy => w10_sub wx (extend5 w4 wy)
- | N10 wx, N6 wy => w10_sub wx (extend4 w5 wy)
- | N10 wx, N7 wy => w10_sub wx (extend3 w6 wy)
- | N10 wx, N8 wy => w10_sub wx (extend2 w7 wy)
- | N10 wx, N9 wy => w10_sub wx (extend1 w8 wy)
- | N10 wx, N10 wy => w10_sub wx wy
- | N10 wx, N11 wy => w11_sub (extend1 w9 wx) wy
- | N10 wx, N12 wy => w12_sub (extend2 w9 wx) wy
- | N10 wx, N13 wy => w13_sub (extend3 w9 wx) wy
- | N10 wx, Nn n wy => subn n (extend n w13 (extend4 w9 wx)) wy
- | N11 wx, N0 wy =>
- if w0_eq0 wy then x
- else w11_sub wx (extend10 w0 (WW w_0 wy))
- | N11 wx, N1 wy => w11_sub wx (extend10 w0 wy)
- | N11 wx, N2 wy => w11_sub wx (extend9 w1 wy)
- | N11 wx, N3 wy => w11_sub wx (extend8 w2 wy)
- | N11 wx, N4 wy => w11_sub wx (extend7 w3 wy)
- | N11 wx, N5 wy => w11_sub wx (extend6 w4 wy)
- | N11 wx, N6 wy => w11_sub wx (extend5 w5 wy)
- | N11 wx, N7 wy => w11_sub wx (extend4 w6 wy)
- | N11 wx, N8 wy => w11_sub wx (extend3 w7 wy)
- | N11 wx, N9 wy => w11_sub wx (extend2 w8 wy)
- | N11 wx, N10 wy => w11_sub wx (extend1 w9 wy)
- | N11 wx, N11 wy => w11_sub wx wy
- | N11 wx, N12 wy => w12_sub (extend1 w10 wx) wy
- | N11 wx, N13 wy => w13_sub (extend2 w10 wx) wy
- | N11 wx, Nn n wy => subn n (extend n w13 (extend3 w10 wx)) wy
- | N12 wx, N0 wy =>
- if w0_eq0 wy then x
- else w12_sub wx (extend11 w0 (WW w_0 wy))
- | N12 wx, N1 wy => w12_sub wx (extend11 w0 wy)
- | N12 wx, N2 wy => w12_sub wx (extend10 w1 wy)
- | N12 wx, N3 wy => w12_sub wx (extend9 w2 wy)
- | N12 wx, N4 wy => w12_sub wx (extend8 w3 wy)
- | N12 wx, N5 wy => w12_sub wx (extend7 w4 wy)
- | N12 wx, N6 wy => w12_sub wx (extend6 w5 wy)
- | N12 wx, N7 wy => w12_sub wx (extend5 w6 wy)
- | N12 wx, N8 wy => w12_sub wx (extend4 w7 wy)
- | N12 wx, N9 wy => w12_sub wx (extend3 w8 wy)
- | N12 wx, N10 wy => w12_sub wx (extend2 w9 wy)
- | N12 wx, N11 wy => w12_sub wx (extend1 w10 wy)
- | N12 wx, N12 wy => w12_sub wx wy
- | N12 wx, N13 wy => w13_sub (extend1 w11 wx) wy
- | N12 wx, Nn n wy => subn n (extend n w13 (extend2 w11 wx)) wy
- | N13 wx, N0 wy =>
- if w0_eq0 wy then x
- else w13_sub wx (extend12 w0 (WW w_0 wy))
- | N13 wx, N1 wy => w13_sub wx (extend12 w0 wy)
- | N13 wx, N2 wy => w13_sub wx (extend11 w1 wy)
- | N13 wx, N3 wy => w13_sub wx (extend10 w2 wy)
- | N13 wx, N4 wy => w13_sub wx (extend9 w3 wy)
- | N13 wx, N5 wy => w13_sub wx (extend8 w4 wy)
- | N13 wx, N6 wy => w13_sub wx (extend7 w5 wy)
- | N13 wx, N7 wy => w13_sub wx (extend6 w6 wy)
- | N13 wx, N8 wy => w13_sub wx (extend5 w7 wy)
- | N13 wx, N9 wy => w13_sub wx (extend4 w8 wy)
- | N13 wx, N10 wy => w13_sub wx (extend3 w9 wy)
- | N13 wx, N11 wy => w13_sub wx (extend2 w10 wy)
- | N13 wx, N12 wy => w13_sub wx (extend1 w11 wy)
- | N13 wx, N13 wy => w13_sub wx wy
- | N13 wx, Nn n wy => subn n (extend n w13 (extend1 w12 wx)) wy
- | Nn n wx, N0 wy =>
- if w0_eq0 wy then x
- else subn n wx (extend n w13 (extend13 w0 (WW w_0 wy)))
- | Nn n wx, N1 wy => subn n wx (extend n w13 (extend13 w0 wy))
- | Nn n wx, N2 wy => subn n wx (extend n w13 (extend12 w1 wy))
- | Nn n wx, N3 wy => subn n wx (extend n w13 (extend11 w2 wy))
- | Nn n wx, N4 wy => subn n wx (extend n w13 (extend10 w3 wy))
- | Nn n wx, N5 wy => subn n wx (extend n w13 (extend9 w4 wy))
- | Nn n wx, N6 wy => subn n wx (extend n w13 (extend8 w5 wy))
- | Nn n wx, N7 wy => subn n wx (extend n w13 (extend7 w6 wy))
- | Nn n wx, N8 wy => subn n wx (extend n w13 (extend6 w7 wy))
- | Nn n wx, N9 wy => subn n wx (extend n w13 (extend5 w8 wy))
- | Nn n wx, N10 wy => subn n wx (extend n w13 (extend4 w9 wy))
- | Nn n wx, N11 wy => subn n wx (extend n w13 (extend3 w10 wy))
- | Nn n wx, N12 wy => subn n wx (extend n w13 (extend2 w11 wy))
- | Nn n wx, N13 wy => subn n wx (extend n w13 (extend1 w12 wy))
- | Nn n wx, Nn m wy =>
- let mn := Max.max n m in
- let d := diff n m in
- subn mn
- (castm (diff_r n m) (extend_tr wx (snd d)))
- (castm (diff_l n m) (extend_tr wy (fst d)))
- end.
+ Definition sub := Eval lazy beta delta [same_level] in
+ (same_level t_ w0_sub w1_sub w2_sub w3_sub w4_sub w5_sub w6_sub w7_sub w8_sub w9_sub w10_sub w11_sub w12_sub w13_sub subn).
+
+ Theorem spec_sub: forall x y, [y] <= [x] -> [sub x y] = [x] - [y].
+ Admitted.
+
+ Theorem spec_sub0: forall x y, [x] < [y] -> [sub x y] = 0.
+ Admitted.
+
+ (***************************************************************)
+ (* *)
+ (* Comparison *)
+ (* *)
+ (***************************************************************)
Definition compare_0 := w0_op.(znz_compare).
Definition comparen_0 :=
@@ -1188,268 +2108,67 @@ Module Make (W0:W0Type).
Definition comparen_13 :=
compare_mn_1 w13 w13 W0 compare_13 (compare_13 W0) compare_13.
- Definition compare x y :=
- match x, y with
- | N0 wx, N0 wy => compare_0 wx wy
- | N0 wx, N1 wy => opp_compare (comparen_0 1 wy wx)
- | N0 wx, N2 wy => opp_compare (comparen_0 2 wy wx)
- | N0 wx, N3 wy => opp_compare (comparen_0 3 wy wx)
- | N0 wx, N4 wy => opp_compare (comparen_0 4 wy wx)
- | N0 wx, N5 wy => opp_compare (comparen_0 5 wy wx)
- | N0 wx, N6 wy => opp_compare (comparen_0 6 wy wx)
- | N0 wx, N7 wy => opp_compare (comparen_0 7 wy wx)
- | N0 wx, N8 wy => opp_compare (comparen_0 8 wy wx)
- | N0 wx, N9 wy => opp_compare (comparen_0 9 wy wx)
- | N0 wx, N10 wy => opp_compare (comparen_0 10 wy wx)
- | N0 wx, N11 wy => opp_compare (comparen_0 11 wy wx)
- | N0 wx, N12 wy => opp_compare (comparen_0 12 wy wx)
- | N0 wx, N13 wy => opp_compare (comparen_0 13 wy wx)
- | N0 wx, Nn n wy =>
- opp_compare (compare_mn_1 w13 w0 w_0 compare_0 (compare_13 W0) (comparen_0 13) (S n) wy wx)
- | N1 wx, N0 wy => comparen_0 1 wx wy
- | N1 wx, N1 wy => compare_1 wx wy
- | N1 wx, N2 wy => opp_compare (comparen_1 1 wy wx)
- | N1 wx, N3 wy => opp_compare (comparen_1 2 wy wx)
- | N1 wx, N4 wy => opp_compare (comparen_1 3 wy wx)
- | N1 wx, N5 wy => opp_compare (comparen_1 4 wy wx)
- | N1 wx, N6 wy => opp_compare (comparen_1 5 wy wx)
- | N1 wx, N7 wy => opp_compare (comparen_1 6 wy wx)
- | N1 wx, N8 wy => opp_compare (comparen_1 7 wy wx)
- | N1 wx, N9 wy => opp_compare (comparen_1 8 wy wx)
- | N1 wx, N10 wy => opp_compare (comparen_1 9 wy wx)
- | N1 wx, N11 wy => opp_compare (comparen_1 10 wy wx)
- | N1 wx, N12 wy => opp_compare (comparen_1 11 wy wx)
- | N1 wx, N13 wy => opp_compare (comparen_1 12 wy wx)
- | N1 wx, Nn n wy =>
- opp_compare (compare_mn_1 w13 w1 W0 compare_1 (compare_13 W0) (comparen_1 12) (S n) wy wx)
- | N2 wx, N0 wy => comparen_0 2 wx wy
- | N2 wx, N1 wy => comparen_1 1 wx wy
- | N2 wx, N2 wy => compare_2 wx wy
- | N2 wx, N3 wy => opp_compare (comparen_2 1 wy wx)
- | N2 wx, N4 wy => opp_compare (comparen_2 2 wy wx)
- | N2 wx, N5 wy => opp_compare (comparen_2 3 wy wx)
- | N2 wx, N6 wy => opp_compare (comparen_2 4 wy wx)
- | N2 wx, N7 wy => opp_compare (comparen_2 5 wy wx)
- | N2 wx, N8 wy => opp_compare (comparen_2 6 wy wx)
- | N2 wx, N9 wy => opp_compare (comparen_2 7 wy wx)
- | N2 wx, N10 wy => opp_compare (comparen_2 8 wy wx)
- | N2 wx, N11 wy => opp_compare (comparen_2 9 wy wx)
- | N2 wx, N12 wy => opp_compare (comparen_2 10 wy wx)
- | N2 wx, N13 wy => opp_compare (comparen_2 11 wy wx)
- | N2 wx, Nn n wy =>
- opp_compare (compare_mn_1 w13 w2 W0 compare_2 (compare_13 W0) (comparen_2 11) (S n) wy wx)
- | N3 wx, N0 wy => comparen_0 3 wx wy
- | N3 wx, N1 wy => comparen_1 2 wx wy
- | N3 wx, N2 wy => comparen_2 1 wx wy
- | N3 wx, N3 wy => compare_3 wx wy
- | N3 wx, N4 wy => opp_compare (comparen_3 1 wy wx)
- | N3 wx, N5 wy => opp_compare (comparen_3 2 wy wx)
- | N3 wx, N6 wy => opp_compare (comparen_3 3 wy wx)
- | N3 wx, N7 wy => opp_compare (comparen_3 4 wy wx)
- | N3 wx, N8 wy => opp_compare (comparen_3 5 wy wx)
- | N3 wx, N9 wy => opp_compare (comparen_3 6 wy wx)
- | N3 wx, N10 wy => opp_compare (comparen_3 7 wy wx)
- | N3 wx, N11 wy => opp_compare (comparen_3 8 wy wx)
- | N3 wx, N12 wy => opp_compare (comparen_3 9 wy wx)
- | N3 wx, N13 wy => opp_compare (comparen_3 10 wy wx)
- | N3 wx, Nn n wy =>
- opp_compare (compare_mn_1 w13 w3 W0 compare_3 (compare_13 W0) (comparen_3 10) (S n) wy wx)
- | N4 wx, N0 wy => comparen_0 4 wx wy
- | N4 wx, N1 wy => comparen_1 3 wx wy
- | N4 wx, N2 wy => comparen_2 2 wx wy
- | N4 wx, N3 wy => comparen_3 1 wx wy
- | N4 wx, N4 wy => compare_4 wx wy
- | N4 wx, N5 wy => opp_compare (comparen_4 1 wy wx)
- | N4 wx, N6 wy => opp_compare (comparen_4 2 wy wx)
- | N4 wx, N7 wy => opp_compare (comparen_4 3 wy wx)
- | N4 wx, N8 wy => opp_compare (comparen_4 4 wy wx)
- | N4 wx, N9 wy => opp_compare (comparen_4 5 wy wx)
- | N4 wx, N10 wy => opp_compare (comparen_4 6 wy wx)
- | N4 wx, N11 wy => opp_compare (comparen_4 7 wy wx)
- | N4 wx, N12 wy => opp_compare (comparen_4 8 wy wx)
- | N4 wx, N13 wy => opp_compare (comparen_4 9 wy wx)
- | N4 wx, Nn n wy =>
- opp_compare (compare_mn_1 w13 w4 W0 compare_4 (compare_13 W0) (comparen_4 9) (S n) wy wx)
- | N5 wx, N0 wy => comparen_0 5 wx wy
- | N5 wx, N1 wy => comparen_1 4 wx wy
- | N5 wx, N2 wy => comparen_2 3 wx wy
- | N5 wx, N3 wy => comparen_3 2 wx wy
- | N5 wx, N4 wy => comparen_4 1 wx wy
- | N5 wx, N5 wy => compare_5 wx wy
- | N5 wx, N6 wy => opp_compare (comparen_5 1 wy wx)
- | N5 wx, N7 wy => opp_compare (comparen_5 2 wy wx)
- | N5 wx, N8 wy => opp_compare (comparen_5 3 wy wx)
- | N5 wx, N9 wy => opp_compare (comparen_5 4 wy wx)
- | N5 wx, N10 wy => opp_compare (comparen_5 5 wy wx)
- | N5 wx, N11 wy => opp_compare (comparen_5 6 wy wx)
- | N5 wx, N12 wy => opp_compare (comparen_5 7 wy wx)
- | N5 wx, N13 wy => opp_compare (comparen_5 8 wy wx)
- | N5 wx, Nn n wy =>
- opp_compare (compare_mn_1 w13 w5 W0 compare_5 (compare_13 W0) (comparen_5 8) (S n) wy wx)
- | N6 wx, N0 wy => comparen_0 6 wx wy
- | N6 wx, N1 wy => comparen_1 5 wx wy
- | N6 wx, N2 wy => comparen_2 4 wx wy
- | N6 wx, N3 wy => comparen_3 3 wx wy
- | N6 wx, N4 wy => comparen_4 2 wx wy
- | N6 wx, N5 wy => comparen_5 1 wx wy
- | N6 wx, N6 wy => compare_6 wx wy
- | N6 wx, N7 wy => opp_compare (comparen_6 1 wy wx)
- | N6 wx, N8 wy => opp_compare (comparen_6 2 wy wx)
- | N6 wx, N9 wy => opp_compare (comparen_6 3 wy wx)
- | N6 wx, N10 wy => opp_compare (comparen_6 4 wy wx)
- | N6 wx, N11 wy => opp_compare (comparen_6 5 wy wx)
- | N6 wx, N12 wy => opp_compare (comparen_6 6 wy wx)
- | N6 wx, N13 wy => opp_compare (comparen_6 7 wy wx)
- | N6 wx, Nn n wy =>
- opp_compare (compare_mn_1 w13 w6 W0 compare_6 (compare_13 W0) (comparen_6 7) (S n) wy wx)
- | N7 wx, N0 wy => comparen_0 7 wx wy
- | N7 wx, N1 wy => comparen_1 6 wx wy
- | N7 wx, N2 wy => comparen_2 5 wx wy
- | N7 wx, N3 wy => comparen_3 4 wx wy
- | N7 wx, N4 wy => comparen_4 3 wx wy
- | N7 wx, N5 wy => comparen_5 2 wx wy
- | N7 wx, N6 wy => comparen_6 1 wx wy
- | N7 wx, N7 wy => compare_7 wx wy
- | N7 wx, N8 wy => opp_compare (comparen_7 1 wy wx)
- | N7 wx, N9 wy => opp_compare (comparen_7 2 wy wx)
- | N7 wx, N10 wy => opp_compare (comparen_7 3 wy wx)
- | N7 wx, N11 wy => opp_compare (comparen_7 4 wy wx)
- | N7 wx, N12 wy => opp_compare (comparen_7 5 wy wx)
- | N7 wx, N13 wy => opp_compare (comparen_7 6 wy wx)
- | N7 wx, Nn n wy =>
- opp_compare (compare_mn_1 w13 w7 W0 compare_7 (compare_13 W0) (comparen_7 6) (S n) wy wx)
- | N8 wx, N0 wy => comparen_0 8 wx wy
- | N8 wx, N1 wy => comparen_1 7 wx wy
- | N8 wx, N2 wy => comparen_2 6 wx wy
- | N8 wx, N3 wy => comparen_3 5 wx wy
- | N8 wx, N4 wy => comparen_4 4 wx wy
- | N8 wx, N5 wy => comparen_5 3 wx wy
- | N8 wx, N6 wy => comparen_6 2 wx wy
- | N8 wx, N7 wy => comparen_7 1 wx wy
- | N8 wx, N8 wy => compare_8 wx wy
- | N8 wx, N9 wy => opp_compare (comparen_8 1 wy wx)
- | N8 wx, N10 wy => opp_compare (comparen_8 2 wy wx)
- | N8 wx, N11 wy => opp_compare (comparen_8 3 wy wx)
- | N8 wx, N12 wy => opp_compare (comparen_8 4 wy wx)
- | N8 wx, N13 wy => opp_compare (comparen_8 5 wy wx)
- | N8 wx, Nn n wy =>
- opp_compare (compare_mn_1 w13 w8 W0 compare_8 (compare_13 W0) (comparen_8 5) (S n) wy wx)
- | N9 wx, N0 wy => comparen_0 9 wx wy
- | N9 wx, N1 wy => comparen_1 8 wx wy
- | N9 wx, N2 wy => comparen_2 7 wx wy
- | N9 wx, N3 wy => comparen_3 6 wx wy
- | N9 wx, N4 wy => comparen_4 5 wx wy
- | N9 wx, N5 wy => comparen_5 4 wx wy
- | N9 wx, N6 wy => comparen_6 3 wx wy
- | N9 wx, N7 wy => comparen_7 2 wx wy
- | N9 wx, N8 wy => comparen_8 1 wx wy
- | N9 wx, N9 wy => compare_9 wx wy
- | N9 wx, N10 wy => opp_compare (comparen_9 1 wy wx)
- | N9 wx, N11 wy => opp_compare (comparen_9 2 wy wx)
- | N9 wx, N12 wy => opp_compare (comparen_9 3 wy wx)
- | N9 wx, N13 wy => opp_compare (comparen_9 4 wy wx)
- | N9 wx, Nn n wy =>
- opp_compare (compare_mn_1 w13 w9 W0 compare_9 (compare_13 W0) (comparen_9 4) (S n) wy wx)
- | N10 wx, N0 wy => comparen_0 10 wx wy
- | N10 wx, N1 wy => comparen_1 9 wx wy
- | N10 wx, N2 wy => comparen_2 8 wx wy
- | N10 wx, N3 wy => comparen_3 7 wx wy
- | N10 wx, N4 wy => comparen_4 6 wx wy
- | N10 wx, N5 wy => comparen_5 5 wx wy
- | N10 wx, N6 wy => comparen_6 4 wx wy
- | N10 wx, N7 wy => comparen_7 3 wx wy
- | N10 wx, N8 wy => comparen_8 2 wx wy
- | N10 wx, N9 wy => comparen_9 1 wx wy
- | N10 wx, N10 wy => compare_10 wx wy
- | N10 wx, N11 wy => opp_compare (comparen_10 1 wy wx)
- | N10 wx, N12 wy => opp_compare (comparen_10 2 wy wx)
- | N10 wx, N13 wy => opp_compare (comparen_10 3 wy wx)
- | N10 wx, Nn n wy =>
- opp_compare (compare_mn_1 w13 w10 W0 compare_10 (compare_13 W0) (comparen_10 3) (S n) wy wx)
- | N11 wx, N0 wy => comparen_0 11 wx wy
- | N11 wx, N1 wy => comparen_1 10 wx wy
- | N11 wx, N2 wy => comparen_2 9 wx wy
- | N11 wx, N3 wy => comparen_3 8 wx wy
- | N11 wx, N4 wy => comparen_4 7 wx wy
- | N11 wx, N5 wy => comparen_5 6 wx wy
- | N11 wx, N6 wy => comparen_6 5 wx wy
- | N11 wx, N7 wy => comparen_7 4 wx wy
- | N11 wx, N8 wy => comparen_8 3 wx wy
- | N11 wx, N9 wy => comparen_9 2 wx wy
- | N11 wx, N10 wy => comparen_10 1 wx wy
- | N11 wx, N11 wy => compare_11 wx wy
- | N11 wx, N12 wy => opp_compare (comparen_11 1 wy wx)
- | N11 wx, N13 wy => opp_compare (comparen_11 2 wy wx)
- | N11 wx, Nn n wy =>
- opp_compare (compare_mn_1 w13 w11 W0 compare_11 (compare_13 W0) (comparen_11 2) (S n) wy wx)
- | N12 wx, N0 wy => comparen_0 12 wx wy
- | N12 wx, N1 wy => comparen_1 11 wx wy
- | N12 wx, N2 wy => comparen_2 10 wx wy
- | N12 wx, N3 wy => comparen_3 9 wx wy
- | N12 wx, N4 wy => comparen_4 8 wx wy
- | N12 wx, N5 wy => comparen_5 7 wx wy
- | N12 wx, N6 wy => comparen_6 6 wx wy
- | N12 wx, N7 wy => comparen_7 5 wx wy
- | N12 wx, N8 wy => comparen_8 4 wx wy
- | N12 wx, N9 wy => comparen_9 3 wx wy
- | N12 wx, N10 wy => comparen_10 2 wx wy
- | N12 wx, N11 wy => comparen_11 1 wx wy
- | N12 wx, N12 wy => compare_12 wx wy
- | N12 wx, N13 wy => opp_compare (comparen_12 1 wy wx)
- | N12 wx, Nn n wy =>
- opp_compare (compare_mn_1 w13 w12 W0 compare_12 (compare_13 W0) (comparen_12 1) (S n) wy wx)
- | N13 wx, N0 wy => comparen_0 13 wx wy
- | N13 wx, N1 wy => comparen_1 12 wx wy
- | N13 wx, N2 wy => comparen_2 11 wx wy
- | N13 wx, N3 wy => comparen_3 10 wx wy
- | N13 wx, N4 wy => comparen_4 9 wx wy
- | N13 wx, N5 wy => comparen_5 8 wx wy
- | N13 wx, N6 wy => comparen_6 7 wx wy
- | N13 wx, N7 wy => comparen_7 6 wx wy
- | N13 wx, N8 wy => comparen_8 5 wx wy
- | N13 wx, N9 wy => comparen_9 4 wx wy
- | N13 wx, N10 wy => comparen_10 3 wx wy
- | N13 wx, N11 wy => comparen_11 2 wx wy
- | N13 wx, N12 wy => comparen_12 1 wx wy
- | N13 wx, N13 wy => compare_13 wx wy
- | N13 wx, Nn n wy =>
- opp_compare (compare_mn_1 w13 w13 W0 compare_13 (compare_13 W0) (comparen_13 0) (S n) wy wx)
- | Nn n wx, N0 wy =>
- compare_mn_1 w13 w0 w_0 compare_0 (compare_13 W0) (comparen_0 13) (S n) wx wy
- | Nn n wx, N1 wy =>
- compare_mn_1 w13 w1 W0 compare_1 (compare_13 W0) (comparen_1 12) (S n) wx wy
- | Nn n wx, N2 wy =>
- compare_mn_1 w13 w2 W0 compare_2 (compare_13 W0) (comparen_2 11) (S n) wx wy
- | Nn n wx, N3 wy =>
- compare_mn_1 w13 w3 W0 compare_3 (compare_13 W0) (comparen_3 10) (S n) wx wy
- | Nn n wx, N4 wy =>
- compare_mn_1 w13 w4 W0 compare_4 (compare_13 W0) (comparen_4 9) (S n) wx wy
- | Nn n wx, N5 wy =>
- compare_mn_1 w13 w5 W0 compare_5 (compare_13 W0) (comparen_5 8) (S n) wx wy
- | Nn n wx, N6 wy =>
- compare_mn_1 w13 w6 W0 compare_6 (compare_13 W0) (comparen_6 7) (S n) wx wy
- | Nn n wx, N7 wy =>
- compare_mn_1 w13 w7 W0 compare_7 (compare_13 W0) (comparen_7 6) (S n) wx wy
- | Nn n wx, N8 wy =>
- compare_mn_1 w13 w8 W0 compare_8 (compare_13 W0) (comparen_8 5) (S n) wx wy
- | Nn n wx, N9 wy =>
- compare_mn_1 w13 w9 W0 compare_9 (compare_13 W0) (comparen_9 4) (S n) wx wy
- | Nn n wx, N10 wy =>
- compare_mn_1 w13 w10 W0 compare_10 (compare_13 W0) (comparen_10 3) (S n) wx wy
- | Nn n wx, N11 wy =>
- compare_mn_1 w13 w11 W0 compare_11 (compare_13 W0) (comparen_11 2) (S n) wx wy
- | Nn n wx, N12 wy =>
- compare_mn_1 w13 w12 W0 compare_12 (compare_13 W0) (comparen_12 1) (S n) wx wy
- | Nn n wx, N13 wy =>
- compare_mn_1 w13 w13 W0 compare_13 (compare_13 W0) (comparen_13 0) (S n) wx wy
- | Nn n wx, Nn m wy =>
+ Definition comparenm n m wx wy :=
let mn := Max.max n m in
let d := diff n m in
let op := make_op mn in
op.(znz_compare)
(castm (diff_r n m) (extend_tr wx (snd d)))
- (castm (diff_l n m) (extend_tr wy (fst d)))
- end.
+ (castm (diff_l n m) (extend_tr wy (fst d))).
+
+ Definition compare := Eval lazy beta delta [iter] in
+ (iter _
+ compare_0
+ (fun n x y => opp_compare (comparen_0 (S n) y x))
+ (fun n => comparen_0 (S n))
+ compare_1
+ (fun n x y => opp_compare (comparen_1 (S n) y x))
+ (fun n => comparen_1 (S n))
+ compare_2
+ (fun n x y => opp_compare (comparen_2 (S n) y x))
+ (fun n => comparen_2 (S n))
+ compare_3
+ (fun n x y => opp_compare (comparen_3 (S n) y x))
+ (fun n => comparen_3 (S n))
+ compare_4
+ (fun n x y => opp_compare (comparen_4 (S n) y x))
+ (fun n => comparen_4 (S n))
+ compare_5
+ (fun n x y => opp_compare (comparen_5 (S n) y x))
+ (fun n => comparen_5 (S n))
+ compare_6
+ (fun n x y => opp_compare (comparen_6 (S n) y x))
+ (fun n => comparen_6 (S n))
+ compare_7
+ (fun n x y => opp_compare (comparen_7 (S n) y x))
+ (fun n => comparen_7 (S n))
+ compare_8
+ (fun n x y => opp_compare (comparen_8 (S n) y x))
+ (fun n => comparen_8 (S n))
+ compare_9
+ (fun n x y => opp_compare (comparen_9 (S n) y x))
+ (fun n => comparen_9 (S n))
+ compare_10
+ (fun n x y => opp_compare (comparen_10 (S n) y x))
+ (fun n => comparen_10 (S n))
+ compare_11
+ (fun n x y => opp_compare (comparen_11 (S n) y x))
+ (fun n => comparen_11 (S n))
+ compare_12
+ (fun n x y => opp_compare (comparen_12 (S n) y x))
+ (fun n => comparen_12 (S n))
+ compare_13
+ (fun n x y => opp_compare (comparen_13 (S n) y x))
+ (fun n => comparen_13 (S n))
+ comparenm).
+
+ Theorem spec_compare: forall x y,
+ match compare x y with
+ Eq => [x] = [y]
+ | Lt => [x] < [y]
+ | Gt => [x] > [y]
+ end.
+ Admitted.
Definition eq_bool x y :=
match compare x y with
@@ -1457,6 +2176,16 @@ Module Make (W0:W0Type).
| _ => false
end.
+ Theorem spec_eq_bool: forall x y,
+ if eq_bool x y then [x] = [y] else [x] <> [y].
+ Admitted.
+
+ (***************************************************************)
+ (* *)
+ (* Multiplication *)
+ (* *)
+ (***************************************************************)
+
Definition w0_mul_c := w0_op.(znz_mul_c).
Definition w1_mul_c := w1_op.(znz_mul_c).
Definition w2_mul_c := w2_op.(znz_mul_c).
@@ -1515,6 +2244,21 @@ Module Make (W0:W0Type).
Eval lazy beta delta [w_mul_add] in
@w_mul_add w13 W0 w13_succ w13_add_c w13_mul_c.
+ Definition w0_0W := w0_op.(znz_0W).
+ Definition w1_0W := w1_op.(znz_0W).
+ Definition w2_0W := w2_op.(znz_0W).
+ Definition w3_0W := w3_op.(znz_0W).
+ Definition w4_0W := w4_op.(znz_0W).
+ Definition w5_0W := w5_op.(znz_0W).
+ Definition w6_0W := w6_op.(znz_0W).
+ Definition w7_0W := w7_op.(znz_0W).
+ Definition w8_0W := w8_op.(znz_0W).
+ Definition w9_0W := w9_op.(znz_0W).
+ Definition w10_0W := w10_op.(znz_0W).
+ Definition w11_0W := w11_op.(znz_0W).
+ Definition w12_0W := w12_op.(znz_0W).
+ Definition w13_0W := w13_op.(znz_0W).
+
Definition w0_mul_add_n1 :=
@gen_mul_add_n1 w0 w_0 w0_op.(znz_WW) w0_0W w0_mul_add.
Definition w1_mul_add_n1 :=
@@ -1544,940 +2288,323 @@ Module Make (W0:W0Type).
Definition w13_mul_add_n1 :=
@gen_mul_add_n1 w13 W0 w13_op.(znz_WW) w13_0W w13_mul_add.
- Definition mul x y :=
- match x, y with
- | N0 wx, N0 wy =>
- reduce_1 (w0_mul_c wx wy)
- | N0 wx, N1 wy =>
- if w0_eq0 wx then zero
- else
- let (w,r) := w0_mul_add_n1 1 wy wx w_0 in
- if w0_eq0 w then N1 r
- else N2 (WW (WW w_0 w) r)
- | N0 wx, N2 wy =>
- if w0_eq0 wx then zero
- else
- let (w,r) := w0_mul_add_n1 2 wy wx w_0 in
- if w0_eq0 w then N2 r
- else N3 (WW (extend1 w0 (WW w_0 w)) r)
- | N0 wx, N3 wy =>
- if w0_eq0 wx then zero
- else
- let (w,r) := w0_mul_add_n1 3 wy wx w_0 in
- if w0_eq0 w then N3 r
- else N4 (WW (extend2 w0 (WW w_0 w)) r)
- | N0 wx, N4 wy =>
- if w0_eq0 wx then zero
- else
- let (w,r) := w0_mul_add_n1 4 wy wx w_0 in
- if w0_eq0 w then N4 r
- else N5 (WW (extend3 w0 (WW w_0 w)) r)
- | N0 wx, N5 wy =>
- if w0_eq0 wx then zero
- else
- let (w,r) := w0_mul_add_n1 5 wy wx w_0 in
- if w0_eq0 w then N5 r
- else N6 (WW (extend4 w0 (WW w_0 w)) r)
- | N0 wx, N6 wy =>
- if w0_eq0 wx then zero
- else
- let (w,r) := w0_mul_add_n1 6 wy wx w_0 in
- if w0_eq0 w then N6 r
- else N7 (WW (extend5 w0 (WW w_0 w)) r)
- | N0 wx, N7 wy =>
- if w0_eq0 wx then zero
- else
- let (w,r) := w0_mul_add_n1 7 wy wx w_0 in
- if w0_eq0 w then N7 r
- else N8 (WW (extend6 w0 (WW w_0 w)) r)
- | N0 wx, N8 wy =>
- if w0_eq0 wx then zero
- else
- let (w,r) := w0_mul_add_n1 8 wy wx w_0 in
- if w0_eq0 w then N8 r
- else N9 (WW (extend7 w0 (WW w_0 w)) r)
- | N0 wx, N9 wy =>
- if w0_eq0 wx then zero
- else
- let (w,r) := w0_mul_add_n1 9 wy wx w_0 in
- if w0_eq0 w then N9 r
- else N10 (WW (extend8 w0 (WW w_0 w)) r)
- | N0 wx, N10 wy =>
- if w0_eq0 wx then zero
- else
- let (w,r) := w0_mul_add_n1 10 wy wx w_0 in
- if w0_eq0 w then N10 r
- else N11 (WW (extend9 w0 (WW w_0 w)) r)
- | N0 wx, N11 wy =>
- if w0_eq0 wx then zero
- else
- let (w,r) := w0_mul_add_n1 11 wy wx w_0 in
- if w0_eq0 w then N11 r
- else N12 (WW (extend10 w0 (WW w_0 w)) r)
- | N0 wx, N12 wy =>
- if w0_eq0 wx then zero
- else
- let (w,r) := w0_mul_add_n1 12 wy wx w_0 in
- if w0_eq0 w then N12 r
- else N13 (WW (extend11 w0 (WW w_0 w)) r)
- | N0 wx, N13 wy =>
- if w0_eq0 wx then zero
- else
- let (w,r) := w0_mul_add_n1 13 wy wx w_0 in
- if w0_eq0 w then N13 r
- else Nn 0 (WW (extend12 w0 (WW w_0 w)) r)
- | N0 wx, Nn n wy =>
- if w0_eq0 wx then zero
- else
- let (w,r) := w13_mul_add_n1 (S n) wy (extend12 w0 (WW w_0 wx)) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | N1 wx, N0 wy =>
- if w0_eq0 wy then zero
- else
- let (w,r) := w0_mul_add_n1 1 wx wy w_0 in
- if w0_eq0 w then N1 r
- else N2 (WW (WW w_0 w) r)
- | N1 wx, N1 wy =>
- N2 (w1_mul_c wx wy)
- | N1 wx, N2 wy =>
- let (w,r) := w1_mul_add_n1 1 wy wx W0 in
- if w1_eq0 w then N2 r
- else N3 (WW (extend1 w0 w) r)
- | N1 wx, N3 wy =>
- let (w,r) := w1_mul_add_n1 2 wy wx W0 in
- if w1_eq0 w then N3 r
- else N4 (WW (extend2 w0 w) r)
- | N1 wx, N4 wy =>
- let (w,r) := w1_mul_add_n1 3 wy wx W0 in
- if w1_eq0 w then N4 r
- else N5 (WW (extend3 w0 w) r)
- | N1 wx, N5 wy =>
- let (w,r) := w1_mul_add_n1 4 wy wx W0 in
- if w1_eq0 w then N5 r
- else N6 (WW (extend4 w0 w) r)
- | N1 wx, N6 wy =>
- let (w,r) := w1_mul_add_n1 5 wy wx W0 in
- if w1_eq0 w then N6 r
- else N7 (WW (extend5 w0 w) r)
- | N1 wx, N7 wy =>
- let (w,r) := w1_mul_add_n1 6 wy wx W0 in
- if w1_eq0 w then N7 r
- else N8 (WW (extend6 w0 w) r)
- | N1 wx, N8 wy =>
- let (w,r) := w1_mul_add_n1 7 wy wx W0 in
- if w1_eq0 w then N8 r
- else N9 (WW (extend7 w0 w) r)
- | N1 wx, N9 wy =>
- let (w,r) := w1_mul_add_n1 8 wy wx W0 in
- if w1_eq0 w then N9 r
- else N10 (WW (extend8 w0 w) r)
- | N1 wx, N10 wy =>
- let (w,r) := w1_mul_add_n1 9 wy wx W0 in
- if w1_eq0 w then N10 r
- else N11 (WW (extend9 w0 w) r)
- | N1 wx, N11 wy =>
- let (w,r) := w1_mul_add_n1 10 wy wx W0 in
- if w1_eq0 w then N11 r
- else N12 (WW (extend10 w0 w) r)
- | N1 wx, N12 wy =>
- let (w,r) := w1_mul_add_n1 11 wy wx W0 in
- if w1_eq0 w then N12 r
- else N13 (WW (extend11 w0 w) r)
- | N1 wx, N13 wy =>
- let (w,r) := w1_mul_add_n1 12 wy wx W0 in
- if w1_eq0 w then N13 r
- else Nn 0 (WW (extend12 w0 w) r)
- | N1 wx, Nn n wy =>
- let (w,r) := w13_mul_add_n1 (S n) wy (extend12 w0 wx) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | N2 wx, N0 wy =>
- if w0_eq0 wy then zero
- else
- let (w,r) := w0_mul_add_n1 2 wx wy w_0 in
- if w0_eq0 w then N2 r
- else N3 (WW (extend1 w0 (WW w_0 w)) r)
- | N2 wx, N1 wy =>
- let (w,r) := w1_mul_add_n1 1 wx wy W0 in
- if w1_eq0 w then N2 r
- else N3 (WW (extend1 w0 w) r)
- | N2 wx, N2 wy =>
- N3 (w2_mul_c wx wy)
- | N2 wx, N3 wy =>
- let (w,r) := w2_mul_add_n1 1 wy wx W0 in
- if w2_eq0 w then N3 r
- else N4 (WW (extend1 w1 w) r)
- | N2 wx, N4 wy =>
- let (w,r) := w2_mul_add_n1 2 wy wx W0 in
- if w2_eq0 w then N4 r
- else N5 (WW (extend2 w1 w) r)
- | N2 wx, N5 wy =>
- let (w,r) := w2_mul_add_n1 3 wy wx W0 in
- if w2_eq0 w then N5 r
- else N6 (WW (extend3 w1 w) r)
- | N2 wx, N6 wy =>
- let (w,r) := w2_mul_add_n1 4 wy wx W0 in
- if w2_eq0 w then N6 r
- else N7 (WW (extend4 w1 w) r)
- | N2 wx, N7 wy =>
- let (w,r) := w2_mul_add_n1 5 wy wx W0 in
- if w2_eq0 w then N7 r
- else N8 (WW (extend5 w1 w) r)
- | N2 wx, N8 wy =>
- let (w,r) := w2_mul_add_n1 6 wy wx W0 in
- if w2_eq0 w then N8 r
- else N9 (WW (extend6 w1 w) r)
- | N2 wx, N9 wy =>
- let (w,r) := w2_mul_add_n1 7 wy wx W0 in
- if w2_eq0 w then N9 r
- else N10 (WW (extend7 w1 w) r)
- | N2 wx, N10 wy =>
- let (w,r) := w2_mul_add_n1 8 wy wx W0 in
- if w2_eq0 w then N10 r
- else N11 (WW (extend8 w1 w) r)
- | N2 wx, N11 wy =>
- let (w,r) := w2_mul_add_n1 9 wy wx W0 in
- if w2_eq0 w then N11 r
- else N12 (WW (extend9 w1 w) r)
- | N2 wx, N12 wy =>
- let (w,r) := w2_mul_add_n1 10 wy wx W0 in
- if w2_eq0 w then N12 r
- else N13 (WW (extend10 w1 w) r)
- | N2 wx, N13 wy =>
- let (w,r) := w2_mul_add_n1 11 wy wx W0 in
- if w2_eq0 w then N13 r
- else Nn 0 (WW (extend11 w1 w) r)
- | N2 wx, Nn n wy =>
- let (w,r) := w13_mul_add_n1 (S n) wy (extend11 w1 wx) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | N3 wx, N0 wy =>
- if w0_eq0 wy then zero
- else
- let (w,r) := w0_mul_add_n1 3 wx wy w_0 in
- if w0_eq0 w then N3 r
- else N4 (WW (extend2 w0 (WW w_0 w)) r)
- | N3 wx, N1 wy =>
- let (w,r) := w1_mul_add_n1 2 wx wy W0 in
- if w1_eq0 w then N3 r
- else N4 (WW (extend2 w0 w) r)
- | N3 wx, N2 wy =>
- let (w,r) := w2_mul_add_n1 1 wx wy W0 in
- if w2_eq0 w then N3 r
- else N4 (WW (extend1 w1 w) r)
- | N3 wx, N3 wy =>
- N4 (w3_mul_c wx wy)
- | N3 wx, N4 wy =>
- let (w,r) := w3_mul_add_n1 1 wy wx W0 in
- if w3_eq0 w then N4 r
- else N5 (WW (extend1 w2 w) r)
- | N3 wx, N5 wy =>
- let (w,r) := w3_mul_add_n1 2 wy wx W0 in
- if w3_eq0 w then N5 r
- else N6 (WW (extend2 w2 w) r)
- | N3 wx, N6 wy =>
- let (w,r) := w3_mul_add_n1 3 wy wx W0 in
- if w3_eq0 w then N6 r
- else N7 (WW (extend3 w2 w) r)
- | N3 wx, N7 wy =>
- let (w,r) := w3_mul_add_n1 4 wy wx W0 in
- if w3_eq0 w then N7 r
- else N8 (WW (extend4 w2 w) r)
- | N3 wx, N8 wy =>
- let (w,r) := w3_mul_add_n1 5 wy wx W0 in
- if w3_eq0 w then N8 r
- else N9 (WW (extend5 w2 w) r)
- | N3 wx, N9 wy =>
- let (w,r) := w3_mul_add_n1 6 wy wx W0 in
- if w3_eq0 w then N9 r
- else N10 (WW (extend6 w2 w) r)
- | N3 wx, N10 wy =>
- let (w,r) := w3_mul_add_n1 7 wy wx W0 in
- if w3_eq0 w then N10 r
- else N11 (WW (extend7 w2 w) r)
- | N3 wx, N11 wy =>
- let (w,r) := w3_mul_add_n1 8 wy wx W0 in
- if w3_eq0 w then N11 r
- else N12 (WW (extend8 w2 w) r)
- | N3 wx, N12 wy =>
- let (w,r) := w3_mul_add_n1 9 wy wx W0 in
- if w3_eq0 w then N12 r
- else N13 (WW (extend9 w2 w) r)
- | N3 wx, N13 wy =>
- let (w,r) := w3_mul_add_n1 10 wy wx W0 in
- if w3_eq0 w then N13 r
- else Nn 0 (WW (extend10 w2 w) r)
- | N3 wx, Nn n wy =>
- let (w,r) := w13_mul_add_n1 (S n) wy (extend10 w2 wx) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | N4 wx, N0 wy =>
- if w0_eq0 wy then zero
- else
- let (w,r) := w0_mul_add_n1 4 wx wy w_0 in
- if w0_eq0 w then N4 r
- else N5 (WW (extend3 w0 (WW w_0 w)) r)
- | N4 wx, N1 wy =>
- let (w,r) := w1_mul_add_n1 3 wx wy W0 in
- if w1_eq0 w then N4 r
- else N5 (WW (extend3 w0 w) r)
- | N4 wx, N2 wy =>
- let (w,r) := w2_mul_add_n1 2 wx wy W0 in
- if w2_eq0 w then N4 r
- else N5 (WW (extend2 w1 w) r)
- | N4 wx, N3 wy =>
- let (w,r) := w3_mul_add_n1 1 wx wy W0 in
- if w3_eq0 w then N4 r
- else N5 (WW (extend1 w2 w) r)
- | N4 wx, N4 wy =>
- N5 (w4_mul_c wx wy)
- | N4 wx, N5 wy =>
- let (w,r) := w4_mul_add_n1 1 wy wx W0 in
- if w4_eq0 w then N5 r
- else N6 (WW (extend1 w3 w) r)
- | N4 wx, N6 wy =>
- let (w,r) := w4_mul_add_n1 2 wy wx W0 in
- if w4_eq0 w then N6 r
- else N7 (WW (extend2 w3 w) r)
- | N4 wx, N7 wy =>
- let (w,r) := w4_mul_add_n1 3 wy wx W0 in
- if w4_eq0 w then N7 r
- else N8 (WW (extend3 w3 w) r)
- | N4 wx, N8 wy =>
- let (w,r) := w4_mul_add_n1 4 wy wx W0 in
- if w4_eq0 w then N8 r
- else N9 (WW (extend4 w3 w) r)
- | N4 wx, N9 wy =>
- let (w,r) := w4_mul_add_n1 5 wy wx W0 in
- if w4_eq0 w then N9 r
- else N10 (WW (extend5 w3 w) r)
- | N4 wx, N10 wy =>
- let (w,r) := w4_mul_add_n1 6 wy wx W0 in
- if w4_eq0 w then N10 r
- else N11 (WW (extend6 w3 w) r)
- | N4 wx, N11 wy =>
- let (w,r) := w4_mul_add_n1 7 wy wx W0 in
- if w4_eq0 w then N11 r
- else N12 (WW (extend7 w3 w) r)
- | N4 wx, N12 wy =>
- let (w,r) := w4_mul_add_n1 8 wy wx W0 in
- if w4_eq0 w then N12 r
- else N13 (WW (extend8 w3 w) r)
- | N4 wx, N13 wy =>
- let (w,r) := w4_mul_add_n1 9 wy wx W0 in
- if w4_eq0 w then N13 r
- else Nn 0 (WW (extend9 w3 w) r)
- | N4 wx, Nn n wy =>
- let (w,r) := w13_mul_add_n1 (S n) wy (extend9 w3 wx) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | N5 wx, N0 wy =>
- if w0_eq0 wy then zero
- else
- let (w,r) := w0_mul_add_n1 5 wx wy w_0 in
- if w0_eq0 w then N5 r
- else N6 (WW (extend4 w0 (WW w_0 w)) r)
- | N5 wx, N1 wy =>
- let (w,r) := w1_mul_add_n1 4 wx wy W0 in
- if w1_eq0 w then N5 r
- else N6 (WW (extend4 w0 w) r)
- | N5 wx, N2 wy =>
- let (w,r) := w2_mul_add_n1 3 wx wy W0 in
- if w2_eq0 w then N5 r
- else N6 (WW (extend3 w1 w) r)
- | N5 wx, N3 wy =>
- let (w,r) := w3_mul_add_n1 2 wx wy W0 in
- if w3_eq0 w then N5 r
- else N6 (WW (extend2 w2 w) r)
- | N5 wx, N4 wy =>
- let (w,r) := w4_mul_add_n1 1 wx wy W0 in
- if w4_eq0 w then N5 r
- else N6 (WW (extend1 w3 w) r)
- | N5 wx, N5 wy =>
- N6 (w5_mul_c wx wy)
- | N5 wx, N6 wy =>
- let (w,r) := w5_mul_add_n1 1 wy wx W0 in
- if w5_eq0 w then N6 r
- else N7 (WW (extend1 w4 w) r)
- | N5 wx, N7 wy =>
- let (w,r) := w5_mul_add_n1 2 wy wx W0 in
- if w5_eq0 w then N7 r
- else N8 (WW (extend2 w4 w) r)
- | N5 wx, N8 wy =>
- let (w,r) := w5_mul_add_n1 3 wy wx W0 in
- if w5_eq0 w then N8 r
- else N9 (WW (extend3 w4 w) r)
- | N5 wx, N9 wy =>
- let (w,r) := w5_mul_add_n1 4 wy wx W0 in
- if w5_eq0 w then N9 r
- else N10 (WW (extend4 w4 w) r)
- | N5 wx, N10 wy =>
- let (w,r) := w5_mul_add_n1 5 wy wx W0 in
- if w5_eq0 w then N10 r
- else N11 (WW (extend5 w4 w) r)
- | N5 wx, N11 wy =>
- let (w,r) := w5_mul_add_n1 6 wy wx W0 in
- if w5_eq0 w then N11 r
- else N12 (WW (extend6 w4 w) r)
- | N5 wx, N12 wy =>
- let (w,r) := w5_mul_add_n1 7 wy wx W0 in
- if w5_eq0 w then N12 r
- else N13 (WW (extend7 w4 w) r)
- | N5 wx, N13 wy =>
- let (w,r) := w5_mul_add_n1 8 wy wx W0 in
- if w5_eq0 w then N13 r
- else Nn 0 (WW (extend8 w4 w) r)
- | N5 wx, Nn n wy =>
- let (w,r) := w13_mul_add_n1 (S n) wy (extend8 w4 wx) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | N6 wx, N0 wy =>
- if w0_eq0 wy then zero
- else
- let (w,r) := w0_mul_add_n1 6 wx wy w_0 in
- if w0_eq0 w then N6 r
- else N7 (WW (extend5 w0 (WW w_0 w)) r)
- | N6 wx, N1 wy =>
- let (w,r) := w1_mul_add_n1 5 wx wy W0 in
- if w1_eq0 w then N6 r
- else N7 (WW (extend5 w0 w) r)
- | N6 wx, N2 wy =>
- let (w,r) := w2_mul_add_n1 4 wx wy W0 in
- if w2_eq0 w then N6 r
- else N7 (WW (extend4 w1 w) r)
- | N6 wx, N3 wy =>
- let (w,r) := w3_mul_add_n1 3 wx wy W0 in
- if w3_eq0 w then N6 r
- else N7 (WW (extend3 w2 w) r)
- | N6 wx, N4 wy =>
- let (w,r) := w4_mul_add_n1 2 wx wy W0 in
- if w4_eq0 w then N6 r
- else N7 (WW (extend2 w3 w) r)
- | N6 wx, N5 wy =>
- let (w,r) := w5_mul_add_n1 1 wx wy W0 in
- if w5_eq0 w then N6 r
- else N7 (WW (extend1 w4 w) r)
- | N6 wx, N6 wy =>
- N7 (w6_mul_c wx wy)
- | N6 wx, N7 wy =>
- let (w,r) := w6_mul_add_n1 1 wy wx W0 in
- if w6_eq0 w then N7 r
- else N8 (WW (extend1 w5 w) r)
- | N6 wx, N8 wy =>
- let (w,r) := w6_mul_add_n1 2 wy wx W0 in
- if w6_eq0 w then N8 r
- else N9 (WW (extend2 w5 w) r)
- | N6 wx, N9 wy =>
- let (w,r) := w6_mul_add_n1 3 wy wx W0 in
- if w6_eq0 w then N9 r
- else N10 (WW (extend3 w5 w) r)
- | N6 wx, N10 wy =>
- let (w,r) := w6_mul_add_n1 4 wy wx W0 in
- if w6_eq0 w then N10 r
- else N11 (WW (extend4 w5 w) r)
- | N6 wx, N11 wy =>
- let (w,r) := w6_mul_add_n1 5 wy wx W0 in
- if w6_eq0 w then N11 r
- else N12 (WW (extend5 w5 w) r)
- | N6 wx, N12 wy =>
- let (w,r) := w6_mul_add_n1 6 wy wx W0 in
- if w6_eq0 w then N12 r
- else N13 (WW (extend6 w5 w) r)
- | N6 wx, N13 wy =>
- let (w,r) := w6_mul_add_n1 7 wy wx W0 in
- if w6_eq0 w then N13 r
- else Nn 0 (WW (extend7 w5 w) r)
- | N6 wx, Nn n wy =>
- let (w,r) := w13_mul_add_n1 (S n) wy (extend7 w5 wx) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | N7 wx, N0 wy =>
- if w0_eq0 wy then zero
- else
- let (w,r) := w0_mul_add_n1 7 wx wy w_0 in
- if w0_eq0 w then N7 r
- else N8 (WW (extend6 w0 (WW w_0 w)) r)
- | N7 wx, N1 wy =>
- let (w,r) := w1_mul_add_n1 6 wx wy W0 in
- if w1_eq0 w then N7 r
- else N8 (WW (extend6 w0 w) r)
- | N7 wx, N2 wy =>
- let (w,r) := w2_mul_add_n1 5 wx wy W0 in
- if w2_eq0 w then N7 r
- else N8 (WW (extend5 w1 w) r)
- | N7 wx, N3 wy =>
- let (w,r) := w3_mul_add_n1 4 wx wy W0 in
- if w3_eq0 w then N7 r
- else N8 (WW (extend4 w2 w) r)
- | N7 wx, N4 wy =>
- let (w,r) := w4_mul_add_n1 3 wx wy W0 in
- if w4_eq0 w then N7 r
- else N8 (WW (extend3 w3 w) r)
- | N7 wx, N5 wy =>
- let (w,r) := w5_mul_add_n1 2 wx wy W0 in
- if w5_eq0 w then N7 r
- else N8 (WW (extend2 w4 w) r)
- | N7 wx, N6 wy =>
- let (w,r) := w6_mul_add_n1 1 wx wy W0 in
- if w6_eq0 w then N7 r
- else N8 (WW (extend1 w5 w) r)
- | N7 wx, N7 wy =>
- N8 (w7_mul_c wx wy)
- | N7 wx, N8 wy =>
- let (w,r) := w7_mul_add_n1 1 wy wx W0 in
- if w7_eq0 w then N8 r
- else N9 (WW (extend1 w6 w) r)
- | N7 wx, N9 wy =>
- let (w,r) := w7_mul_add_n1 2 wy wx W0 in
- if w7_eq0 w then N9 r
- else N10 (WW (extend2 w6 w) r)
- | N7 wx, N10 wy =>
- let (w,r) := w7_mul_add_n1 3 wy wx W0 in
- if w7_eq0 w then N10 r
- else N11 (WW (extend3 w6 w) r)
- | N7 wx, N11 wy =>
- let (w,r) := w7_mul_add_n1 4 wy wx W0 in
- if w7_eq0 w then N11 r
- else N12 (WW (extend4 w6 w) r)
- | N7 wx, N12 wy =>
- let (w,r) := w7_mul_add_n1 5 wy wx W0 in
- if w7_eq0 w then N12 r
- else N13 (WW (extend5 w6 w) r)
- | N7 wx, N13 wy =>
- let (w,r) := w7_mul_add_n1 6 wy wx W0 in
- if w7_eq0 w then N13 r
- else Nn 0 (WW (extend6 w6 w) r)
- | N7 wx, Nn n wy =>
- let (w,r) := w13_mul_add_n1 (S n) wy (extend6 w6 wx) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | N8 wx, N0 wy =>
- if w0_eq0 wy then zero
- else
- let (w,r) := w0_mul_add_n1 8 wx wy w_0 in
- if w0_eq0 w then N8 r
- else N9 (WW (extend7 w0 (WW w_0 w)) r)
- | N8 wx, N1 wy =>
- let (w,r) := w1_mul_add_n1 7 wx wy W0 in
- if w1_eq0 w then N8 r
- else N9 (WW (extend7 w0 w) r)
- | N8 wx, N2 wy =>
- let (w,r) := w2_mul_add_n1 6 wx wy W0 in
- if w2_eq0 w then N8 r
- else N9 (WW (extend6 w1 w) r)
- | N8 wx, N3 wy =>
- let (w,r) := w3_mul_add_n1 5 wx wy W0 in
- if w3_eq0 w then N8 r
- else N9 (WW (extend5 w2 w) r)
- | N8 wx, N4 wy =>
- let (w,r) := w4_mul_add_n1 4 wx wy W0 in
- if w4_eq0 w then N8 r
- else N9 (WW (extend4 w3 w) r)
- | N8 wx, N5 wy =>
- let (w,r) := w5_mul_add_n1 3 wx wy W0 in
- if w5_eq0 w then N8 r
- else N9 (WW (extend3 w4 w) r)
- | N8 wx, N6 wy =>
- let (w,r) := w6_mul_add_n1 2 wx wy W0 in
- if w6_eq0 w then N8 r
- else N9 (WW (extend2 w5 w) r)
- | N8 wx, N7 wy =>
- let (w,r) := w7_mul_add_n1 1 wx wy W0 in
- if w7_eq0 w then N8 r
- else N9 (WW (extend1 w6 w) r)
- | N8 wx, N8 wy =>
- N9 (w8_mul_c wx wy)
- | N8 wx, N9 wy =>
- let (w,r) := w8_mul_add_n1 1 wy wx W0 in
- if w8_eq0 w then N9 r
- else N10 (WW (extend1 w7 w) r)
- | N8 wx, N10 wy =>
- let (w,r) := w8_mul_add_n1 2 wy wx W0 in
- if w8_eq0 w then N10 r
- else N11 (WW (extend2 w7 w) r)
- | N8 wx, N11 wy =>
- let (w,r) := w8_mul_add_n1 3 wy wx W0 in
- if w8_eq0 w then N11 r
- else N12 (WW (extend3 w7 w) r)
- | N8 wx, N12 wy =>
- let (w,r) := w8_mul_add_n1 4 wy wx W0 in
- if w8_eq0 w then N12 r
- else N13 (WW (extend4 w7 w) r)
- | N8 wx, N13 wy =>
- let (w,r) := w8_mul_add_n1 5 wy wx W0 in
- if w8_eq0 w then N13 r
- else Nn 0 (WW (extend5 w7 w) r)
- | N8 wx, Nn n wy =>
- let (w,r) := w13_mul_add_n1 (S n) wy (extend5 w7 wx) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | N9 wx, N0 wy =>
- if w0_eq0 wy then zero
- else
- let (w,r) := w0_mul_add_n1 9 wx wy w_0 in
- if w0_eq0 w then N9 r
- else N10 (WW (extend8 w0 (WW w_0 w)) r)
- | N9 wx, N1 wy =>
- let (w,r) := w1_mul_add_n1 8 wx wy W0 in
- if w1_eq0 w then N9 r
- else N10 (WW (extend8 w0 w) r)
- | N9 wx, N2 wy =>
- let (w,r) := w2_mul_add_n1 7 wx wy W0 in
- if w2_eq0 w then N9 r
- else N10 (WW (extend7 w1 w) r)
- | N9 wx, N3 wy =>
- let (w,r) := w3_mul_add_n1 6 wx wy W0 in
- if w3_eq0 w then N9 r
- else N10 (WW (extend6 w2 w) r)
- | N9 wx, N4 wy =>
- let (w,r) := w4_mul_add_n1 5 wx wy W0 in
- if w4_eq0 w then N9 r
- else N10 (WW (extend5 w3 w) r)
- | N9 wx, N5 wy =>
- let (w,r) := w5_mul_add_n1 4 wx wy W0 in
- if w5_eq0 w then N9 r
- else N10 (WW (extend4 w4 w) r)
- | N9 wx, N6 wy =>
- let (w,r) := w6_mul_add_n1 3 wx wy W0 in
- if w6_eq0 w then N9 r
- else N10 (WW (extend3 w5 w) r)
- | N9 wx, N7 wy =>
- let (w,r) := w7_mul_add_n1 2 wx wy W0 in
- if w7_eq0 w then N9 r
- else N10 (WW (extend2 w6 w) r)
- | N9 wx, N8 wy =>
- let (w,r) := w8_mul_add_n1 1 wx wy W0 in
- if w8_eq0 w then N9 r
- else N10 (WW (extend1 w7 w) r)
- | N9 wx, N9 wy =>
- N10 (w9_mul_c wx wy)
- | N9 wx, N10 wy =>
- let (w,r) := w9_mul_add_n1 1 wy wx W0 in
- if w9_eq0 w then N10 r
- else N11 (WW (extend1 w8 w) r)
- | N9 wx, N11 wy =>
- let (w,r) := w9_mul_add_n1 2 wy wx W0 in
- if w9_eq0 w then N11 r
- else N12 (WW (extend2 w8 w) r)
- | N9 wx, N12 wy =>
- let (w,r) := w9_mul_add_n1 3 wy wx W0 in
- if w9_eq0 w then N12 r
- else N13 (WW (extend3 w8 w) r)
- | N9 wx, N13 wy =>
- let (w,r) := w9_mul_add_n1 4 wy wx W0 in
- if w9_eq0 w then N13 r
- else Nn 0 (WW (extend4 w8 w) r)
- | N9 wx, Nn n wy =>
- let (w,r) := w13_mul_add_n1 (S n) wy (extend4 w8 wx) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | N10 wx, N0 wy =>
- if w0_eq0 wy then zero
- else
- let (w,r) := w0_mul_add_n1 10 wx wy w_0 in
- if w0_eq0 w then N10 r
- else N11 (WW (extend9 w0 (WW w_0 w)) r)
- | N10 wx, N1 wy =>
- let (w,r) := w1_mul_add_n1 9 wx wy W0 in
- if w1_eq0 w then N10 r
- else N11 (WW (extend9 w0 w) r)
- | N10 wx, N2 wy =>
- let (w,r) := w2_mul_add_n1 8 wx wy W0 in
- if w2_eq0 w then N10 r
- else N11 (WW (extend8 w1 w) r)
- | N10 wx, N3 wy =>
- let (w,r) := w3_mul_add_n1 7 wx wy W0 in
- if w3_eq0 w then N10 r
- else N11 (WW (extend7 w2 w) r)
- | N10 wx, N4 wy =>
- let (w,r) := w4_mul_add_n1 6 wx wy W0 in
- if w4_eq0 w then N10 r
- else N11 (WW (extend6 w3 w) r)
- | N10 wx, N5 wy =>
- let (w,r) := w5_mul_add_n1 5 wx wy W0 in
- if w5_eq0 w then N10 r
- else N11 (WW (extend5 w4 w) r)
- | N10 wx, N6 wy =>
- let (w,r) := w6_mul_add_n1 4 wx wy W0 in
- if w6_eq0 w then N10 r
- else N11 (WW (extend4 w5 w) r)
- | N10 wx, N7 wy =>
- let (w,r) := w7_mul_add_n1 3 wx wy W0 in
- if w7_eq0 w then N10 r
- else N11 (WW (extend3 w6 w) r)
- | N10 wx, N8 wy =>
- let (w,r) := w8_mul_add_n1 2 wx wy W0 in
- if w8_eq0 w then N10 r
- else N11 (WW (extend2 w7 w) r)
- | N10 wx, N9 wy =>
- let (w,r) := w9_mul_add_n1 1 wx wy W0 in
- if w9_eq0 w then N10 r
- else N11 (WW (extend1 w8 w) r)
- | N10 wx, N10 wy =>
- N11 (w10_mul_c wx wy)
- | N10 wx, N11 wy =>
- let (w,r) := w10_mul_add_n1 1 wy wx W0 in
- if w10_eq0 w then N11 r
- else N12 (WW (extend1 w9 w) r)
- | N10 wx, N12 wy =>
- let (w,r) := w10_mul_add_n1 2 wy wx W0 in
- if w10_eq0 w then N12 r
- else N13 (WW (extend2 w9 w) r)
- | N10 wx, N13 wy =>
- let (w,r) := w10_mul_add_n1 3 wy wx W0 in
- if w10_eq0 w then N13 r
- else Nn 0 (WW (extend3 w9 w) r)
- | N10 wx, Nn n wy =>
- let (w,r) := w13_mul_add_n1 (S n) wy (extend3 w9 wx) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | N11 wx, N0 wy =>
- if w0_eq0 wy then zero
- else
- let (w,r) := w0_mul_add_n1 11 wx wy w_0 in
- if w0_eq0 w then N11 r
- else N12 (WW (extend10 w0 (WW w_0 w)) r)
- | N11 wx, N1 wy =>
- let (w,r) := w1_mul_add_n1 10 wx wy W0 in
- if w1_eq0 w then N11 r
- else N12 (WW (extend10 w0 w) r)
- | N11 wx, N2 wy =>
- let (w,r) := w2_mul_add_n1 9 wx wy W0 in
- if w2_eq0 w then N11 r
- else N12 (WW (extend9 w1 w) r)
- | N11 wx, N3 wy =>
- let (w,r) := w3_mul_add_n1 8 wx wy W0 in
- if w3_eq0 w then N11 r
- else N12 (WW (extend8 w2 w) r)
- | N11 wx, N4 wy =>
- let (w,r) := w4_mul_add_n1 7 wx wy W0 in
- if w4_eq0 w then N11 r
- else N12 (WW (extend7 w3 w) r)
- | N11 wx, N5 wy =>
- let (w,r) := w5_mul_add_n1 6 wx wy W0 in
- if w5_eq0 w then N11 r
- else N12 (WW (extend6 w4 w) r)
- | N11 wx, N6 wy =>
- let (w,r) := w6_mul_add_n1 5 wx wy W0 in
- if w6_eq0 w then N11 r
- else N12 (WW (extend5 w5 w) r)
- | N11 wx, N7 wy =>
- let (w,r) := w7_mul_add_n1 4 wx wy W0 in
- if w7_eq0 w then N11 r
- else N12 (WW (extend4 w6 w) r)
- | N11 wx, N8 wy =>
- let (w,r) := w8_mul_add_n1 3 wx wy W0 in
- if w8_eq0 w then N11 r
- else N12 (WW (extend3 w7 w) r)
- | N11 wx, N9 wy =>
- let (w,r) := w9_mul_add_n1 2 wx wy W0 in
- if w9_eq0 w then N11 r
- else N12 (WW (extend2 w8 w) r)
- | N11 wx, N10 wy =>
- let (w,r) := w10_mul_add_n1 1 wx wy W0 in
- if w10_eq0 w then N11 r
- else N12 (WW (extend1 w9 w) r)
- | N11 wx, N11 wy =>
- N12 (w11_mul_c wx wy)
- | N11 wx, N12 wy =>
- let (w,r) := w11_mul_add_n1 1 wy wx W0 in
- if w11_eq0 w then N12 r
- else N13 (WW (extend1 w10 w) r)
- | N11 wx, N13 wy =>
- let (w,r) := w11_mul_add_n1 2 wy wx W0 in
- if w11_eq0 w then N13 r
- else Nn 0 (WW (extend2 w10 w) r)
- | N11 wx, Nn n wy =>
- let (w,r) := w13_mul_add_n1 (S n) wy (extend2 w10 wx) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | N12 wx, N0 wy =>
- if w0_eq0 wy then zero
- else
- let (w,r) := w0_mul_add_n1 12 wx wy w_0 in
- if w0_eq0 w then N12 r
- else N13 (WW (extend11 w0 (WW w_0 w)) r)
- | N12 wx, N1 wy =>
- let (w,r) := w1_mul_add_n1 11 wx wy W0 in
- if w1_eq0 w then N12 r
- else N13 (WW (extend11 w0 w) r)
- | N12 wx, N2 wy =>
- let (w,r) := w2_mul_add_n1 10 wx wy W0 in
- if w2_eq0 w then N12 r
- else N13 (WW (extend10 w1 w) r)
- | N12 wx, N3 wy =>
- let (w,r) := w3_mul_add_n1 9 wx wy W0 in
- if w3_eq0 w then N12 r
- else N13 (WW (extend9 w2 w) r)
- | N12 wx, N4 wy =>
- let (w,r) := w4_mul_add_n1 8 wx wy W0 in
- if w4_eq0 w then N12 r
- else N13 (WW (extend8 w3 w) r)
- | N12 wx, N5 wy =>
- let (w,r) := w5_mul_add_n1 7 wx wy W0 in
- if w5_eq0 w then N12 r
- else N13 (WW (extend7 w4 w) r)
- | N12 wx, N6 wy =>
- let (w,r) := w6_mul_add_n1 6 wx wy W0 in
- if w6_eq0 w then N12 r
- else N13 (WW (extend6 w5 w) r)
- | N12 wx, N7 wy =>
- let (w,r) := w7_mul_add_n1 5 wx wy W0 in
- if w7_eq0 w then N12 r
- else N13 (WW (extend5 w6 w) r)
- | N12 wx, N8 wy =>
- let (w,r) := w8_mul_add_n1 4 wx wy W0 in
- if w8_eq0 w then N12 r
- else N13 (WW (extend4 w7 w) r)
- | N12 wx, N9 wy =>
- let (w,r) := w9_mul_add_n1 3 wx wy W0 in
- if w9_eq0 w then N12 r
- else N13 (WW (extend3 w8 w) r)
- | N12 wx, N10 wy =>
- let (w,r) := w10_mul_add_n1 2 wx wy W0 in
- if w10_eq0 w then N12 r
- else N13 (WW (extend2 w9 w) r)
- | N12 wx, N11 wy =>
- let (w,r) := w11_mul_add_n1 1 wx wy W0 in
- if w11_eq0 w then N12 r
- else N13 (WW (extend1 w10 w) r)
- | N12 wx, N12 wy =>
- N13 (w12_mul_c wx wy)
- | N12 wx, N13 wy =>
- let (w,r) := w12_mul_add_n1 1 wy wx W0 in
- if w12_eq0 w then N13 r
- else Nn 0 (WW (extend1 w11 w) r)
- | N12 wx, Nn n wy =>
- let (w,r) := w13_mul_add_n1 (S n) wy (extend1 w11 wx) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | N13 wx, N0 wy =>
- if w0_eq0 wy then zero
- else
- let (w,r) := w0_mul_add_n1 13 wx wy w_0 in
- if w0_eq0 w then N13 r
- else Nn 0 (WW (extend12 w0 (WW w_0 w)) r)
- | N13 wx, N1 wy =>
- let (w,r) := w1_mul_add_n1 12 wx wy W0 in
- if w1_eq0 w then N13 r
- else Nn 0 (WW (extend12 w0 w) r)
- | N13 wx, N2 wy =>
- let (w,r) := w2_mul_add_n1 11 wx wy W0 in
- if w2_eq0 w then N13 r
- else Nn 0 (WW (extend11 w1 w) r)
- | N13 wx, N3 wy =>
- let (w,r) := w3_mul_add_n1 10 wx wy W0 in
- if w3_eq0 w then N13 r
- else Nn 0 (WW (extend10 w2 w) r)
- | N13 wx, N4 wy =>
- let (w,r) := w4_mul_add_n1 9 wx wy W0 in
- if w4_eq0 w then N13 r
- else Nn 0 (WW (extend9 w3 w) r)
- | N13 wx, N5 wy =>
- let (w,r) := w5_mul_add_n1 8 wx wy W0 in
- if w5_eq0 w then N13 r
- else Nn 0 (WW (extend8 w4 w) r)
- | N13 wx, N6 wy =>
- let (w,r) := w6_mul_add_n1 7 wx wy W0 in
- if w6_eq0 w then N13 r
- else Nn 0 (WW (extend7 w5 w) r)
- | N13 wx, N7 wy =>
- let (w,r) := w7_mul_add_n1 6 wx wy W0 in
- if w7_eq0 w then N13 r
- else Nn 0 (WW (extend6 w6 w) r)
- | N13 wx, N8 wy =>
- let (w,r) := w8_mul_add_n1 5 wx wy W0 in
- if w8_eq0 w then N13 r
- else Nn 0 (WW (extend5 w7 w) r)
- | N13 wx, N9 wy =>
- let (w,r) := w9_mul_add_n1 4 wx wy W0 in
- if w9_eq0 w then N13 r
- else Nn 0 (WW (extend4 w8 w) r)
- | N13 wx, N10 wy =>
- let (w,r) := w10_mul_add_n1 3 wx wy W0 in
- if w10_eq0 w then N13 r
- else Nn 0 (WW (extend3 w9 w) r)
- | N13 wx, N11 wy =>
- let (w,r) := w11_mul_add_n1 2 wx wy W0 in
- if w11_eq0 w then N13 r
- else Nn 0 (WW (extend2 w10 w) r)
- | N13 wx, N12 wy =>
- let (w,r) := w12_mul_add_n1 1 wx wy W0 in
- if w12_eq0 w then N13 r
- else Nn 0 (WW (extend1 w11 w) r)
- | N13 wx, N13 wy =>
- Nn 0 (w13_mul_c wx wy)
- | N13 wx, Nn n wy =>
- let (w,r) := w13_mul_add_n1 (S n) wy wx W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | Nn n wx, N0 wy =>
- if w0_eq0 wy then zero
- else
- let (w,r) := w13_mul_add_n1 (S n) wx (extend12 w0 (WW w_0 wy)) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | Nn n wx, N1 wy =>
- let (w,r) := w13_mul_add_n1 (S n) wx (extend12 w0 wy) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | Nn n wx, N2 wy =>
- let (w,r) := w13_mul_add_n1 (S n) wx (extend11 w1 wy) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | Nn n wx, N3 wy =>
- let (w,r) := w13_mul_add_n1 (S n) wx (extend10 w2 wy) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | Nn n wx, N4 wy =>
- let (w,r) := w13_mul_add_n1 (S n) wx (extend9 w3 wy) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | Nn n wx, N5 wy =>
- let (w,r) := w13_mul_add_n1 (S n) wx (extend8 w4 wy) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | Nn n wx, N6 wy =>
- let (w,r) := w13_mul_add_n1 (S n) wx (extend7 w5 wy) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | Nn n wx, N7 wy =>
- let (w,r) := w13_mul_add_n1 (S n) wx (extend6 w6 wy) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | Nn n wx, N8 wy =>
- let (w,r) := w13_mul_add_n1 (S n) wx (extend5 w7 wy) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | Nn n wx, N9 wy =>
- let (w,r) := w13_mul_add_n1 (S n) wx (extend4 w8 wy) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | Nn n wx, N10 wy =>
- let (w,r) := w13_mul_add_n1 (S n) wx (extend3 w9 wy) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | Nn n wx, N11 wy =>
- let (w,r) := w13_mul_add_n1 (S n) wx (extend2 w10 wy) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | Nn n wx, N12 wy =>
- let (w,r) := w13_mul_add_n1 (S n) wx (extend1 w11 wy) W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | Nn n wx, N13 wy =>
- let (w,r) := w13_mul_add_n1 (S n) wx wy W0 in
- if w13_eq0 w then Nn n r
- else Nn (S n) (WW (extend n w13 (WW W0 w)) r)
- | Nn n wx, Nn m wy =>
+ Let to_Z0 n :=
+ match n return word w0 (S n) -> t_ with
+ | 0%nat => fun x => N1 x
+ | 1%nat => fun x => N2 x
+ | 2%nat => fun x => N3 x
+ | 3%nat => fun x => N4 x
+ | 4%nat => fun x => N5 x
+ | 5%nat => fun x => N6 x
+ | 6%nat => fun x => N7 x
+ | 7%nat => fun x => N8 x
+ | 8%nat => fun x => N9 x
+ | 9%nat => fun x => N10 x
+ | 10%nat => fun x => N11 x
+ | 11%nat => fun x => N12 x
+ | 12%nat => fun x => N13 x
+ | 13%nat => fun x => Nn 0 x
+ | 14%nat => fun x => Nn 1 x
+ | _ => fun _ => N0 w_0
+ end.
+
+ Let to_Z1 n :=
+ match n return word w1 (S n) -> t_ with
+ | 0%nat => fun x => N2 x
+ | 1%nat => fun x => N3 x
+ | 2%nat => fun x => N4 x
+ | 3%nat => fun x => N5 x
+ | 4%nat => fun x => N6 x
+ | 5%nat => fun x => N7 x
+ | 6%nat => fun x => N8 x
+ | 7%nat => fun x => N9 x
+ | 8%nat => fun x => N10 x
+ | 9%nat => fun x => N11 x
+ | 10%nat => fun x => N12 x
+ | 11%nat => fun x => N13 x
+ | 12%nat => fun x => Nn 0 x
+ | 13%nat => fun x => Nn 1 x
+ | _ => fun _ => N0 w_0
+ end.
+
+ Let to_Z2 n :=
+ match n return word w2 (S n) -> t_ with
+ | 0%nat => fun x => N3 x
+ | 1%nat => fun x => N4 x
+ | 2%nat => fun x => N5 x
+ | 3%nat => fun x => N6 x
+ | 4%nat => fun x => N7 x
+ | 5%nat => fun x => N8 x
+ | 6%nat => fun x => N9 x
+ | 7%nat => fun x => N10 x
+ | 8%nat => fun x => N11 x
+ | 9%nat => fun x => N12 x
+ | 10%nat => fun x => N13 x
+ | 11%nat => fun x => Nn 0 x
+ | 12%nat => fun x => Nn 1 x
+ | _ => fun _ => N0 w_0
+ end.
+
+ Let to_Z3 n :=
+ match n return word w3 (S n) -> t_ with
+ | 0%nat => fun x => N4 x
+ | 1%nat => fun x => N5 x
+ | 2%nat => fun x => N6 x
+ | 3%nat => fun x => N7 x
+ | 4%nat => fun x => N8 x
+ | 5%nat => fun x => N9 x
+ | 6%nat => fun x => N10 x
+ | 7%nat => fun x => N11 x
+ | 8%nat => fun x => N12 x
+ | 9%nat => fun x => N13 x
+ | 10%nat => fun x => Nn 0 x
+ | 11%nat => fun x => Nn 1 x
+ | _ => fun _ => N0 w_0
+ end.
+
+ Let to_Z4 n :=
+ match n return word w4 (S n) -> t_ with
+ | 0%nat => fun x => N5 x
+ | 1%nat => fun x => N6 x
+ | 2%nat => fun x => N7 x
+ | 3%nat => fun x => N8 x
+ | 4%nat => fun x => N9 x
+ | 5%nat => fun x => N10 x
+ | 6%nat => fun x => N11 x
+ | 7%nat => fun x => N12 x
+ | 8%nat => fun x => N13 x
+ | 9%nat => fun x => Nn 0 x
+ | 10%nat => fun x => Nn 1 x
+ | _ => fun _ => N0 w_0
+ end.
+
+ Let to_Z5 n :=
+ match n return word w5 (S n) -> t_ with
+ | 0%nat => fun x => N6 x
+ | 1%nat => fun x => N7 x
+ | 2%nat => fun x => N8 x
+ | 3%nat => fun x => N9 x
+ | 4%nat => fun x => N10 x
+ | 5%nat => fun x => N11 x
+ | 6%nat => fun x => N12 x
+ | 7%nat => fun x => N13 x
+ | 8%nat => fun x => Nn 0 x
+ | 9%nat => fun x => Nn 1 x
+ | _ => fun _ => N0 w_0
+ end.
+
+ Let to_Z6 n :=
+ match n return word w6 (S n) -> t_ with
+ | 0%nat => fun x => N7 x
+ | 1%nat => fun x => N8 x
+ | 2%nat => fun x => N9 x
+ | 3%nat => fun x => N10 x
+ | 4%nat => fun x => N11 x
+ | 5%nat => fun x => N12 x
+ | 6%nat => fun x => N13 x
+ | 7%nat => fun x => Nn 0 x
+ | 8%nat => fun x => Nn 1 x
+ | _ => fun _ => N0 w_0
+ end.
+
+ Let to_Z7 n :=
+ match n return word w7 (S n) -> t_ with
+ | 0%nat => fun x => N8 x
+ | 1%nat => fun x => N9 x
+ | 2%nat => fun x => N10 x
+ | 3%nat => fun x => N11 x
+ | 4%nat => fun x => N12 x
+ | 5%nat => fun x => N13 x
+ | 6%nat => fun x => Nn 0 x
+ | 7%nat => fun x => Nn 1 x
+ | _ => fun _ => N0 w_0
+ end.
+
+ Let to_Z8 n :=
+ match n return word w8 (S n) -> t_ with
+ | 0%nat => fun x => N9 x
+ | 1%nat => fun x => N10 x
+ | 2%nat => fun x => N11 x
+ | 3%nat => fun x => N12 x
+ | 4%nat => fun x => N13 x
+ | 5%nat => fun x => Nn 0 x
+ | 6%nat => fun x => Nn 1 x
+ | _ => fun _ => N0 w_0
+ end.
+
+ Let to_Z9 n :=
+ match n return word w9 (S n) -> t_ with
+ | 0%nat => fun x => N10 x
+ | 1%nat => fun x => N11 x
+ | 2%nat => fun x => N12 x
+ | 3%nat => fun x => N13 x
+ | 4%nat => fun x => Nn 0 x
+ | 5%nat => fun x => Nn 1 x
+ | _ => fun _ => N0 w_0
+ end.
+
+ Let to_Z10 n :=
+ match n return word w10 (S n) -> t_ with
+ | 0%nat => fun x => N11 x
+ | 1%nat => fun x => N12 x
+ | 2%nat => fun x => N13 x
+ | 3%nat => fun x => Nn 0 x
+ | 4%nat => fun x => Nn 1 x
+ | _ => fun _ => N0 w_0
+ end.
+
+ Let to_Z11 n :=
+ match n return word w11 (S n) -> t_ with
+ | 0%nat => fun x => N12 x
+ | 1%nat => fun x => N13 x
+ | 2%nat => fun x => Nn 0 x
+ | 3%nat => fun x => Nn 1 x
+ | _ => fun _ => N0 w_0
+ end.
+
+ Let to_Z12 n :=
+ match n return word w12 (S n) -> t_ with
+ | 0%nat => fun x => N13 x
+ | 1%nat => fun x => Nn 0 x
+ | 2%nat => fun x => Nn 1 x
+ | _ => fun _ => N0 w_0
+ end.
+
+ Definition w0_mul n x y :=
+ let (w,r) := w0_mul_add_n1 (S n) x y w_0 in
+ if w0_eq0 w then to_Z0 n r
+ else to_Z0 (S n) (WW (extend0 n w) r).
+
+ Definition w1_mul n x y :=
+ let (w,r) := w1_mul_add_n1 (S n) x y W0 in
+ if w1_eq0 w then to_Z1 n r
+ else to_Z1 (S n) (WW (extend1 n w) r).
+
+ Definition w2_mul n x y :=
+ let (w,r) := w2_mul_add_n1 (S n) x y W0 in
+ if w2_eq0 w then to_Z2 n r
+ else to_Z2 (S n) (WW (extend2 n w) r).
+
+ Definition w3_mul n x y :=
+ let (w,r) := w3_mul_add_n1 (S n) x y W0 in
+ if w3_eq0 w then to_Z3 n r
+ else to_Z3 (S n) (WW (extend3 n w) r).
+
+ Definition w4_mul n x y :=
+ let (w,r) := w4_mul_add_n1 (S n) x y W0 in
+ if w4_eq0 w then to_Z4 n r
+ else to_Z4 (S n) (WW (extend4 n w) r).
+
+ Definition w5_mul n x y :=
+ let (w,r) := w5_mul_add_n1 (S n) x y W0 in
+ if w5_eq0 w then to_Z5 n r
+ else to_Z5 (S n) (WW (extend5 n w) r).
+
+ Definition w6_mul n x y :=
+ let (w,r) := w6_mul_add_n1 (S n) x y W0 in
+ if w6_eq0 w then to_Z6 n r
+ else to_Z6 (S n) (WW (extend6 n w) r).
+
+ Definition w7_mul n x y :=
+ let (w,r) := w7_mul_add_n1 (S n) x y W0 in
+ if w7_eq0 w then to_Z7 n r
+ else to_Z7 (S n) (WW (extend7 n w) r).
+
+ Definition w8_mul n x y :=
+ let (w,r) := w8_mul_add_n1 (S n) x y W0 in
+ if w8_eq0 w then to_Z8 n r
+ else to_Z8 (S n) (WW (extend8 n w) r).
+
+ Definition w9_mul n x y :=
+ let (w,r) := w9_mul_add_n1 (S n) x y W0 in
+ if w9_eq0 w then to_Z9 n r
+ else to_Z9 (S n) (WW (extend9 n w) r).
+
+ Definition w10_mul n x y :=
+ let (w,r) := w10_mul_add_n1 (S n) x y W0 in
+ if w10_eq0 w then to_Z10 n r
+ else to_Z10 (S n) (WW (extend10 n w) r).
+
+ Definition w11_mul n x y :=
+ let (w,r) := w11_mul_add_n1 (S n) x y W0 in
+ if w11_eq0 w then to_Z11 n r
+ else to_Z11 (S n) (WW (extend11 n w) r).
+
+ Definition w12_mul n x y :=
+ let (w,r) := w12_mul_add_n1 (S n) x y W0 in
+ if w12_eq0 w then to_Z12 n r
+ else to_Z12 (S n) (WW (extend12 n w) r).
+
+ Definition w13_mul n x y :=
+ let (w,r) := w13_mul_add_n1 (S n) x y W0 in
+ if w13_eq0 w then Nn n r
+ else Nn (S n) (WW (extend13 n w) r).
+
+ Definition mulnm n m x y :=
let mn := Max.max n m in
let d := diff n m in
let op := make_op mn in
reduce_n (S mn) (op.(znz_mul_c)
- (castm (diff_r n m) (extend_tr wx (snd d)))
- (castm (diff_l n m) (extend_tr wy (fst d))))
- end.
+ (castm (diff_r n m) (extend_tr x (snd d)))
+ (castm (diff_l n m) (extend_tr y (fst d)))).
+
+ Definition mul := Eval lazy beta delta [iter0] in
+ (iter0 t_
+ (fun x y => reduce_1 (w0_mul_c x y))
+ (fun n x y => w0_mul n y x)
+ w0_mul
+ (fun x y => reduce_2 (w1_mul_c x y))
+ (fun n x y => w1_mul n y x)
+ w1_mul
+ (fun x y => reduce_3 (w2_mul_c x y))
+ (fun n x y => w2_mul n y x)
+ w2_mul
+ (fun x y => reduce_4 (w3_mul_c x y))
+ (fun n x y => w3_mul n y x)
+ w3_mul
+ (fun x y => reduce_5 (w4_mul_c x y))
+ (fun n x y => w4_mul n y x)
+ w4_mul
+ (fun x y => reduce_6 (w5_mul_c x y))
+ (fun n x y => w5_mul n y x)
+ w5_mul
+ (fun x y => reduce_7 (w6_mul_c x y))
+ (fun n x y => w6_mul n y x)
+ w6_mul
+ (fun x y => reduce_8 (w7_mul_c x y))
+ (fun n x y => w7_mul n y x)
+ w7_mul
+ (fun x y => reduce_9 (w8_mul_c x y))
+ (fun n x y => w8_mul n y x)
+ w8_mul
+ (fun x y => reduce_10 (w9_mul_c x y))
+ (fun n x y => w9_mul n y x)
+ w9_mul
+ (fun x y => reduce_11 (w10_mul_c x y))
+ (fun n x y => w10_mul n y x)
+ w10_mul
+ (fun x y => reduce_12 (w11_mul_c x y))
+ (fun n x y => w11_mul n y x)
+ w11_mul
+ (fun x y => reduce_13 (w12_mul_c x y))
+ (fun n x y => w12_mul n y x)
+ w12_mul
+ (fun x y => reduce_14 (w13_mul_c x y))
+ (fun n x y => w13_mul n y x)
+ w13_mul
+ mulnm
+ (fun _ => N0 w_0)
+ (fun _ => N0 w_0)
+ ).
+
+ Theorem spec_mul: forall x y, [mul x y] = [x] * [y].
+ Admitted.
+
+ (***************************************************************)
+ (* *)
+ (* Square *)
+ (* *)
+ (***************************************************************)
Definition w0_square_c := w0_op.(znz_square_c).
Definition w1_square_c := w1_op.(znz_square_c).
@@ -2515,6 +2642,15 @@ Module Make (W0:W0Type).
Nn (S n) (op.(znz_square_c) wx)
end.
+ Theorem spec_square: forall x, [square x] = [x] * [x].
+Admitted.
+
+ (***************************************************************)
+ (* *)
+ (* Power *)
+ (* *)
+ (***************************************************************)
+
Fixpoint power_pos (x:t) (p:positive) {struct p} : t :=
match p with
| xH => x
@@ -2522,6 +2658,15 @@ Module Make (W0:W0Type).
| xI p => mul (square (power_pos x p)) x
end.
+ Theorem spec_power_pos: forall x n, [power_pos x n] = [x] ^ Zpos n.
+ Admitted.
+
+ (***************************************************************)
+ (* *)
+ (* Square root *)
+ (* *)
+ (***************************************************************)
+
Definition w0_sqrt := w0_op.(znz_sqrt).
Definition w1_sqrt := w1_op.(znz_sqrt).
Definition w2_sqrt := w2_op.(znz_sqrt).
@@ -2558,6 +2703,15 @@ Module Make (W0:W0Type).
reduce_n n (op.(znz_sqrt) wx)
end.
+ Theorem spec_sqrt: forall x, [sqrt x] ^ 2 <= [x] < ([sqrt x] + 1) ^ 2.
+Admitted.
+
+ (***************************************************************)
+ (* *)
+ (* Division *)
+ (* *)
+ (***************************************************************)
+
Definition w0_div_gt := w0_op.(znz_div_gt).
Definition w1_div_gt := w1_op.(znz_div_gt).
Definition w2_div_gt := w2_op.(znz_div_gt).
@@ -2573,669 +2727,180 @@ Module Make (W0:W0Type).
Definition w12_div_gt := w12_op.(znz_div_gt).
Definition w13_div_gt := w13_op.(znz_div_gt).
- Definition w0_divn1 :=
+ Definition w0_divn1 n x y :=
+ let (u, v) :=
gen_divn1 w0_op.(znz_zdigits) w0_op.(znz_0)
w0_op.(znz_WW) w0_op.(znz_head0)
w0_op.(znz_add_mul_div) w0_op.(znz_div21)
- w0_op.(znz_compare) w0_op.(znz_sub).
- Definition w1_divn1 :=
+ w0_op.(znz_compare) w0_op.(znz_sub) (S n) x y in
+ (to_Z0 _ u, N0 v).
+ Definition w1_divn1 n x y :=
+ let (u, v) :=
gen_divn1 w1_op.(znz_zdigits) w1_op.(znz_0)
w1_op.(znz_WW) w1_op.(znz_head0)
w1_op.(znz_add_mul_div) w1_op.(znz_div21)
- w1_op.(znz_compare) w1_op.(znz_sub).
- Definition w2_divn1 :=
+ w1_op.(znz_compare) w1_op.(znz_sub) (S n) x y in
+ (to_Z1 _ u, N1 v).
+ Definition w2_divn1 n x y :=
+ let (u, v) :=
gen_divn1 w2_op.(znz_zdigits) w2_op.(znz_0)
w2_op.(znz_WW) w2_op.(znz_head0)
w2_op.(znz_add_mul_div) w2_op.(znz_div21)
- w2_op.(znz_compare) w2_op.(znz_sub).
- Definition w3_divn1 :=
+ w2_op.(znz_compare) w2_op.(znz_sub) (S n) x y in
+ (to_Z2 _ u, N2 v).
+ Definition w3_divn1 n x y :=
+ let (u, v) :=
gen_divn1 w3_op.(znz_zdigits) w3_op.(znz_0)
w3_op.(znz_WW) w3_op.(znz_head0)
w3_op.(znz_add_mul_div) w3_op.(znz_div21)
- w3_op.(znz_compare) w3_op.(znz_sub).
- Definition w4_divn1 :=
+ w3_op.(znz_compare) w3_op.(znz_sub) (S n) x y in
+ (to_Z3 _ u, N3 v).
+ Definition w4_divn1 n x y :=
+ let (u, v) :=
gen_divn1 w4_op.(znz_zdigits) w4_op.(znz_0)
w4_op.(znz_WW) w4_op.(znz_head0)
w4_op.(znz_add_mul_div) w4_op.(znz_div21)
- w4_op.(znz_compare) w4_op.(znz_sub).
- Definition w5_divn1 :=
+ w4_op.(znz_compare) w4_op.(znz_sub) (S n) x y in
+ (to_Z4 _ u, N4 v).
+ Definition w5_divn1 n x y :=
+ let (u, v) :=
gen_divn1 w5_op.(znz_zdigits) w5_op.(znz_0)
w5_op.(znz_WW) w5_op.(znz_head0)
w5_op.(znz_add_mul_div) w5_op.(znz_div21)
- w5_op.(znz_compare) w5_op.(znz_sub).
- Definition w6_divn1 :=
+ w5_op.(znz_compare) w5_op.(znz_sub) (S n) x y in
+ (to_Z5 _ u, N5 v).
+ Definition w6_divn1 n x y :=
+ let (u, v) :=
gen_divn1 w6_op.(znz_zdigits) w6_op.(znz_0)
w6_op.(znz_WW) w6_op.(znz_head0)
w6_op.(znz_add_mul_div) w6_op.(znz_div21)
- w6_op.(znz_compare) w6_op.(znz_sub).
- Definition w7_divn1 :=
+ w6_op.(znz_compare) w6_op.(znz_sub) (S n) x y in
+ (to_Z6 _ u, N6 v).
+ Definition w7_divn1 n x y :=
+ let (u, v) :=
gen_divn1 w7_op.(znz_zdigits) w7_op.(znz_0)
w7_op.(znz_WW) w7_op.(znz_head0)
w7_op.(znz_add_mul_div) w7_op.(znz_div21)
- w7_op.(znz_compare) w7_op.(znz_sub).
- Definition w8_divn1 :=
+ w7_op.(znz_compare) w7_op.(znz_sub) (S n) x y in
+ (to_Z7 _ u, N7 v).
+ Definition w8_divn1 n x y :=
+ let (u, v) :=
gen_divn1 w8_op.(znz_zdigits) w8_op.(znz_0)
w8_op.(znz_WW) w8_op.(znz_head0)
w8_op.(znz_add_mul_div) w8_op.(znz_div21)
- w8_op.(znz_compare) w8_op.(znz_sub).
- Definition w9_divn1 :=
+ w8_op.(znz_compare) w8_op.(znz_sub) (S n) x y in
+ (to_Z8 _ u, N8 v).
+ Definition w9_divn1 n x y :=
+ let (u, v) :=
gen_divn1 w9_op.(znz_zdigits) w9_op.(znz_0)
w9_op.(znz_WW) w9_op.(znz_head0)
w9_op.(znz_add_mul_div) w9_op.(znz_div21)
- w9_op.(znz_compare) w9_op.(znz_sub).
- Definition w10_divn1 :=
+ w9_op.(znz_compare) w9_op.(znz_sub) (S n) x y in
+ (to_Z9 _ u, N9 v).
+ Definition w10_divn1 n x y :=
+ let (u, v) :=
gen_divn1 w10_op.(znz_zdigits) w10_op.(znz_0)
w10_op.(znz_WW) w10_op.(znz_head0)
w10_op.(znz_add_mul_div) w10_op.(znz_div21)
- w10_op.(znz_compare) w10_op.(znz_sub).
- Definition w11_divn1 :=
+ w10_op.(znz_compare) w10_op.(znz_sub) (S n) x y in
+ (to_Z10 _ u, N10 v).
+ Definition w11_divn1 n x y :=
+ let (u, v) :=
gen_divn1 w11_op.(znz_zdigits) w11_op.(znz_0)
w11_op.(znz_WW) w11_op.(znz_head0)
w11_op.(znz_add_mul_div) w11_op.(znz_div21)
- w11_op.(znz_compare) w11_op.(znz_sub).
- Definition w12_divn1 :=
+ w11_op.(znz_compare) w11_op.(znz_sub) (S n) x y in
+ (to_Z11 _ u, N11 v).
+ Definition w12_divn1 n x y :=
+ let (u, v) :=
gen_divn1 w12_op.(znz_zdigits) w12_op.(znz_0)
w12_op.(znz_WW) w12_op.(znz_head0)
w12_op.(znz_add_mul_div) w12_op.(znz_div21)
- w12_op.(znz_compare) w12_op.(znz_sub).
- Definition w13_divn1 :=
+ w12_op.(znz_compare) w12_op.(znz_sub) (S n) x y in
+ (to_Z12 _ u, N12 v).
+ Definition w13_divn1 n x y :=
+ let (u, v) :=
gen_divn1 w13_op.(znz_zdigits) w13_op.(znz_0)
w13_op.(znz_WW) w13_op.(znz_head0)
w13_op.(znz_add_mul_div) w13_op.(znz_div21)
- w13_op.(znz_compare) w13_op.(znz_sub).
-
- Definition div_gt x y :=
- match x, y with
- | N0 wx, N0 wy => let (q, r):= w0_div_gt wx wy in (reduce_0 q, reduce_0 r)
- | N0 wx, N1 wy =>
- let wx':= GenBase.extend w0_0W 0 wx in
- let (q, r):= w1_div_gt wx' wy in
- (reduce_1 q, reduce_1 r)
- | N0 wx, N2 wy =>
- let wx':= GenBase.extend w0_0W 1 wx in
- let (q, r):= w2_div_gt wx' wy in
- (reduce_2 q, reduce_2 r)
- | N0 wx, N3 wy =>
- let wx':= GenBase.extend w0_0W 2 wx in
- let (q, r):= w3_div_gt wx' wy in
- (reduce_3 q, reduce_3 r)
- | N0 wx, N4 wy =>
- let wx':= GenBase.extend w0_0W 3 wx in
- let (q, r):= w4_div_gt wx' wy in
- (reduce_4 q, reduce_4 r)
- | N0 wx, N5 wy =>
- let wx':= GenBase.extend w0_0W 4 wx in
- let (q, r):= w5_div_gt wx' wy in
- (reduce_5 q, reduce_5 r)
- | N0 wx, N6 wy =>
- let wx':= GenBase.extend w0_0W 5 wx in
- let (q, r):= w6_div_gt wx' wy in
- (reduce_6 q, reduce_6 r)
- | N0 wx, N7 wy =>
- let wx':= GenBase.extend w0_0W 6 wx in
- let (q, r):= w7_div_gt wx' wy in
- (reduce_7 q, reduce_7 r)
- | N0 wx, N8 wy =>
- let wx':= GenBase.extend w0_0W 7 wx in
- let (q, r):= w8_div_gt wx' wy in
- (reduce_8 q, reduce_8 r)
- | N0 wx, N9 wy =>
- let wx':= GenBase.extend w0_0W 8 wx in
- let (q, r):= w9_div_gt wx' wy in
- (reduce_9 q, reduce_9 r)
- | N0 wx, N10 wy =>
- let wx':= GenBase.extend w0_0W 9 wx in
- let (q, r):= w10_div_gt wx' wy in
- (reduce_10 q, reduce_10 r)
- | N0 wx, N11 wy =>
- let wx':= GenBase.extend w0_0W 10 wx in
- let (q, r):= w11_div_gt wx' wy in
- (reduce_11 q, reduce_11 r)
- | N0 wx, N12 wy =>
- let wx':= GenBase.extend w0_0W 11 wx in
- let (q, r):= w12_div_gt wx' wy in
- (reduce_12 q, reduce_12 r)
- | N0 wx, N13 wy =>
- let wx':= GenBase.extend w0_0W 12 wx in
- let (q, r):= w13_div_gt wx' wy in
- (reduce_13 q, reduce_13 r)
- | N0 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w0_0W 13 wx) in
- let (q, r):= (make_op n).(znz_div_gt) wx' wy in
- (reduce_n n q, reduce_n n r)
- | N1 wx, N0 wy => let (q, r):= w0_divn1 1 wx wy in (reduce_1 q, reduce_0 r)
- | N1 wx, N1 wy => let (q, r):= w1_div_gt wx wy in (reduce_1 q, reduce_1 r)
- | N1 wx, N2 wy =>
- let wx':= GenBase.extend w1_0W 0 wx in
- let (q, r):= w2_div_gt wx' wy in
- (reduce_2 q, reduce_2 r)
- | N1 wx, N3 wy =>
- let wx':= GenBase.extend w1_0W 1 wx in
- let (q, r):= w3_div_gt wx' wy in
- (reduce_3 q, reduce_3 r)
- | N1 wx, N4 wy =>
- let wx':= GenBase.extend w1_0W 2 wx in
- let (q, r):= w4_div_gt wx' wy in
- (reduce_4 q, reduce_4 r)
- | N1 wx, N5 wy =>
- let wx':= GenBase.extend w1_0W 3 wx in
- let (q, r):= w5_div_gt wx' wy in
- (reduce_5 q, reduce_5 r)
- | N1 wx, N6 wy =>
- let wx':= GenBase.extend w1_0W 4 wx in
- let (q, r):= w6_div_gt wx' wy in
- (reduce_6 q, reduce_6 r)
- | N1 wx, N7 wy =>
- let wx':= GenBase.extend w1_0W 5 wx in
- let (q, r):= w7_div_gt wx' wy in
- (reduce_7 q, reduce_7 r)
- | N1 wx, N8 wy =>
- let wx':= GenBase.extend w1_0W 6 wx in
- let (q, r):= w8_div_gt wx' wy in
- (reduce_8 q, reduce_8 r)
- | N1 wx, N9 wy =>
- let wx':= GenBase.extend w1_0W 7 wx in
- let (q, r):= w9_div_gt wx' wy in
- (reduce_9 q, reduce_9 r)
- | N1 wx, N10 wy =>
- let wx':= GenBase.extend w1_0W 8 wx in
- let (q, r):= w10_div_gt wx' wy in
- (reduce_10 q, reduce_10 r)
- | N1 wx, N11 wy =>
- let wx':= GenBase.extend w1_0W 9 wx in
- let (q, r):= w11_div_gt wx' wy in
- (reduce_11 q, reduce_11 r)
- | N1 wx, N12 wy =>
- let wx':= GenBase.extend w1_0W 10 wx in
- let (q, r):= w12_div_gt wx' wy in
- (reduce_12 q, reduce_12 r)
- | N1 wx, N13 wy =>
- let wx':= GenBase.extend w1_0W 11 wx in
- let (q, r):= w13_div_gt wx' wy in
- (reduce_13 q, reduce_13 r)
- | N1 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w1_0W 12 wx) in
- let (q, r):= (make_op n).(znz_div_gt) wx' wy in
- (reduce_n n q, reduce_n n r)
- | N2 wx, N0 wy => let (q, r):= w0_divn1 2 wx wy in (reduce_2 q, reduce_0 r)
- | N2 wx, N1 wy => let (q, r):= w1_divn1 1 wx wy in (reduce_2 q, reduce_1 r)
- | N2 wx, N2 wy => let (q, r):= w2_div_gt wx wy in (reduce_2 q, reduce_2 r)
- | N2 wx, N3 wy =>
- let wx':= GenBase.extend w2_0W 0 wx in
- let (q, r):= w3_div_gt wx' wy in
- (reduce_3 q, reduce_3 r)
- | N2 wx, N4 wy =>
- let wx':= GenBase.extend w2_0W 1 wx in
- let (q, r):= w4_div_gt wx' wy in
- (reduce_4 q, reduce_4 r)
- | N2 wx, N5 wy =>
- let wx':= GenBase.extend w2_0W 2 wx in
- let (q, r):= w5_div_gt wx' wy in
- (reduce_5 q, reduce_5 r)
- | N2 wx, N6 wy =>
- let wx':= GenBase.extend w2_0W 3 wx in
- let (q, r):= w6_div_gt wx' wy in
- (reduce_6 q, reduce_6 r)
- | N2 wx, N7 wy =>
- let wx':= GenBase.extend w2_0W 4 wx in
- let (q, r):= w7_div_gt wx' wy in
- (reduce_7 q, reduce_7 r)
- | N2 wx, N8 wy =>
- let wx':= GenBase.extend w2_0W 5 wx in
- let (q, r):= w8_div_gt wx' wy in
- (reduce_8 q, reduce_8 r)
- | N2 wx, N9 wy =>
- let wx':= GenBase.extend w2_0W 6 wx in
- let (q, r):= w9_div_gt wx' wy in
- (reduce_9 q, reduce_9 r)
- | N2 wx, N10 wy =>
- let wx':= GenBase.extend w2_0W 7 wx in
- let (q, r):= w10_div_gt wx' wy in
- (reduce_10 q, reduce_10 r)
- | N2 wx, N11 wy =>
- let wx':= GenBase.extend w2_0W 8 wx in
- let (q, r):= w11_div_gt wx' wy in
- (reduce_11 q, reduce_11 r)
- | N2 wx, N12 wy =>
- let wx':= GenBase.extend w2_0W 9 wx in
- let (q, r):= w12_div_gt wx' wy in
- (reduce_12 q, reduce_12 r)
- | N2 wx, N13 wy =>
- let wx':= GenBase.extend w2_0W 10 wx in
- let (q, r):= w13_div_gt wx' wy in
- (reduce_13 q, reduce_13 r)
- | N2 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w2_0W 11 wx) in
- let (q, r):= (make_op n).(znz_div_gt) wx' wy in
- (reduce_n n q, reduce_n n r)
- | N3 wx, N0 wy => let (q, r):= w0_divn1 3 wx wy in (reduce_3 q, reduce_0 r)
- | N3 wx, N1 wy => let (q, r):= w1_divn1 2 wx wy in (reduce_3 q, reduce_1 r)
- | N3 wx, N2 wy => let (q, r):= w2_divn1 1 wx wy in (reduce_3 q, reduce_2 r)
- | N3 wx, N3 wy => let (q, r):= w3_div_gt wx wy in (reduce_3 q, reduce_3 r)
- | N3 wx, N4 wy =>
- let wx':= GenBase.extend w3_0W 0 wx in
- let (q, r):= w4_div_gt wx' wy in
- (reduce_4 q, reduce_4 r)
- | N3 wx, N5 wy =>
- let wx':= GenBase.extend w3_0W 1 wx in
- let (q, r):= w5_div_gt wx' wy in
- (reduce_5 q, reduce_5 r)
- | N3 wx, N6 wy =>
- let wx':= GenBase.extend w3_0W 2 wx in
- let (q, r):= w6_div_gt wx' wy in
- (reduce_6 q, reduce_6 r)
- | N3 wx, N7 wy =>
- let wx':= GenBase.extend w3_0W 3 wx in
- let (q, r):= w7_div_gt wx' wy in
- (reduce_7 q, reduce_7 r)
- | N3 wx, N8 wy =>
- let wx':= GenBase.extend w3_0W 4 wx in
- let (q, r):= w8_div_gt wx' wy in
- (reduce_8 q, reduce_8 r)
- | N3 wx, N9 wy =>
- let wx':= GenBase.extend w3_0W 5 wx in
- let (q, r):= w9_div_gt wx' wy in
- (reduce_9 q, reduce_9 r)
- | N3 wx, N10 wy =>
- let wx':= GenBase.extend w3_0W 6 wx in
- let (q, r):= w10_div_gt wx' wy in
- (reduce_10 q, reduce_10 r)
- | N3 wx, N11 wy =>
- let wx':= GenBase.extend w3_0W 7 wx in
- let (q, r):= w11_div_gt wx' wy in
- (reduce_11 q, reduce_11 r)
- | N3 wx, N12 wy =>
- let wx':= GenBase.extend w3_0W 8 wx in
- let (q, r):= w12_div_gt wx' wy in
- (reduce_12 q, reduce_12 r)
- | N3 wx, N13 wy =>
- let wx':= GenBase.extend w3_0W 9 wx in
- let (q, r):= w13_div_gt wx' wy in
- (reduce_13 q, reduce_13 r)
- | N3 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w3_0W 10 wx) in
- let (q, r):= (make_op n).(znz_div_gt) wx' wy in
- (reduce_n n q, reduce_n n r)
- | N4 wx, N0 wy => let (q, r):= w0_divn1 4 wx wy in (reduce_4 q, reduce_0 r)
- | N4 wx, N1 wy => let (q, r):= w1_divn1 3 wx wy in (reduce_4 q, reduce_1 r)
- | N4 wx, N2 wy => let (q, r):= w2_divn1 2 wx wy in (reduce_4 q, reduce_2 r)
- | N4 wx, N3 wy => let (q, r):= w3_divn1 1 wx wy in (reduce_4 q, reduce_3 r)
- | N4 wx, N4 wy => let (q, r):= w4_div_gt wx wy in (reduce_4 q, reduce_4 r)
- | N4 wx, N5 wy =>
- let wx':= GenBase.extend w4_0W 0 wx in
- let (q, r):= w5_div_gt wx' wy in
- (reduce_5 q, reduce_5 r)
- | N4 wx, N6 wy =>
- let wx':= GenBase.extend w4_0W 1 wx in
- let (q, r):= w6_div_gt wx' wy in
- (reduce_6 q, reduce_6 r)
- | N4 wx, N7 wy =>
- let wx':= GenBase.extend w4_0W 2 wx in
- let (q, r):= w7_div_gt wx' wy in
- (reduce_7 q, reduce_7 r)
- | N4 wx, N8 wy =>
- let wx':= GenBase.extend w4_0W 3 wx in
- let (q, r):= w8_div_gt wx' wy in
- (reduce_8 q, reduce_8 r)
- | N4 wx, N9 wy =>
- let wx':= GenBase.extend w4_0W 4 wx in
- let (q, r):= w9_div_gt wx' wy in
- (reduce_9 q, reduce_9 r)
- | N4 wx, N10 wy =>
- let wx':= GenBase.extend w4_0W 5 wx in
- let (q, r):= w10_div_gt wx' wy in
- (reduce_10 q, reduce_10 r)
- | N4 wx, N11 wy =>
- let wx':= GenBase.extend w4_0W 6 wx in
- let (q, r):= w11_div_gt wx' wy in
- (reduce_11 q, reduce_11 r)
- | N4 wx, N12 wy =>
- let wx':= GenBase.extend w4_0W 7 wx in
- let (q, r):= w12_div_gt wx' wy in
- (reduce_12 q, reduce_12 r)
- | N4 wx, N13 wy =>
- let wx':= GenBase.extend w4_0W 8 wx in
- let (q, r):= w13_div_gt wx' wy in
- (reduce_13 q, reduce_13 r)
- | N4 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w4_0W 9 wx) in
- let (q, r):= (make_op n).(znz_div_gt) wx' wy in
- (reduce_n n q, reduce_n n r)
- | N5 wx, N0 wy => let (q, r):= w0_divn1 5 wx wy in (reduce_5 q, reduce_0 r)
- | N5 wx, N1 wy => let (q, r):= w1_divn1 4 wx wy in (reduce_5 q, reduce_1 r)
- | N5 wx, N2 wy => let (q, r):= w2_divn1 3 wx wy in (reduce_5 q, reduce_2 r)
- | N5 wx, N3 wy => let (q, r):= w3_divn1 2 wx wy in (reduce_5 q, reduce_3 r)
- | N5 wx, N4 wy => let (q, r):= w4_divn1 1 wx wy in (reduce_5 q, reduce_4 r)
- | N5 wx, N5 wy => let (q, r):= w5_div_gt wx wy in (reduce_5 q, reduce_5 r)
- | N5 wx, N6 wy =>
- let wx':= GenBase.extend w5_0W 0 wx in
- let (q, r):= w6_div_gt wx' wy in
- (reduce_6 q, reduce_6 r)
- | N5 wx, N7 wy =>
- let wx':= GenBase.extend w5_0W 1 wx in
- let (q, r):= w7_div_gt wx' wy in
- (reduce_7 q, reduce_7 r)
- | N5 wx, N8 wy =>
- let wx':= GenBase.extend w5_0W 2 wx in
- let (q, r):= w8_div_gt wx' wy in
- (reduce_8 q, reduce_8 r)
- | N5 wx, N9 wy =>
- let wx':= GenBase.extend w5_0W 3 wx in
- let (q, r):= w9_div_gt wx' wy in
- (reduce_9 q, reduce_9 r)
- | N5 wx, N10 wy =>
- let wx':= GenBase.extend w5_0W 4 wx in
- let (q, r):= w10_div_gt wx' wy in
- (reduce_10 q, reduce_10 r)
- | N5 wx, N11 wy =>
- let wx':= GenBase.extend w5_0W 5 wx in
- let (q, r):= w11_div_gt wx' wy in
- (reduce_11 q, reduce_11 r)
- | N5 wx, N12 wy =>
- let wx':= GenBase.extend w5_0W 6 wx in
- let (q, r):= w12_div_gt wx' wy in
- (reduce_12 q, reduce_12 r)
- | N5 wx, N13 wy =>
- let wx':= GenBase.extend w5_0W 7 wx in
- let (q, r):= w13_div_gt wx' wy in
- (reduce_13 q, reduce_13 r)
- | N5 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w5_0W 8 wx) in
- let (q, r):= (make_op n).(znz_div_gt) wx' wy in
- (reduce_n n q, reduce_n n r)
- | N6 wx, N0 wy => let (q, r):= w0_divn1 6 wx wy in (reduce_6 q, reduce_0 r)
- | N6 wx, N1 wy => let (q, r):= w1_divn1 5 wx wy in (reduce_6 q, reduce_1 r)
- | N6 wx, N2 wy => let (q, r):= w2_divn1 4 wx wy in (reduce_6 q, reduce_2 r)
- | N6 wx, N3 wy => let (q, r):= w3_divn1 3 wx wy in (reduce_6 q, reduce_3 r)
- | N6 wx, N4 wy => let (q, r):= w4_divn1 2 wx wy in (reduce_6 q, reduce_4 r)
- | N6 wx, N5 wy => let (q, r):= w5_divn1 1 wx wy in (reduce_6 q, reduce_5 r)
- | N6 wx, N6 wy => let (q, r):= w6_div_gt wx wy in (reduce_6 q, reduce_6 r)
- | N6 wx, N7 wy =>
- let wx':= GenBase.extend w6_0W 0 wx in
- let (q, r):= w7_div_gt wx' wy in
- (reduce_7 q, reduce_7 r)
- | N6 wx, N8 wy =>
- let wx':= GenBase.extend w6_0W 1 wx in
- let (q, r):= w8_div_gt wx' wy in
- (reduce_8 q, reduce_8 r)
- | N6 wx, N9 wy =>
- let wx':= GenBase.extend w6_0W 2 wx in
- let (q, r):= w9_div_gt wx' wy in
- (reduce_9 q, reduce_9 r)
- | N6 wx, N10 wy =>
- let wx':= GenBase.extend w6_0W 3 wx in
- let (q, r):= w10_div_gt wx' wy in
- (reduce_10 q, reduce_10 r)
- | N6 wx, N11 wy =>
- let wx':= GenBase.extend w6_0W 4 wx in
- let (q, r):= w11_div_gt wx' wy in
- (reduce_11 q, reduce_11 r)
- | N6 wx, N12 wy =>
- let wx':= GenBase.extend w6_0W 5 wx in
- let (q, r):= w12_div_gt wx' wy in
- (reduce_12 q, reduce_12 r)
- | N6 wx, N13 wy =>
- let wx':= GenBase.extend w6_0W 6 wx in
- let (q, r):= w13_div_gt wx' wy in
- (reduce_13 q, reduce_13 r)
- | N6 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w6_0W 7 wx) in
- let (q, r):= (make_op n).(znz_div_gt) wx' wy in
- (reduce_n n q, reduce_n n r)
- | N7 wx, N0 wy => let (q, r):= w0_divn1 7 wx wy in (reduce_7 q, reduce_0 r)
- | N7 wx, N1 wy => let (q, r):= w1_divn1 6 wx wy in (reduce_7 q, reduce_1 r)
- | N7 wx, N2 wy => let (q, r):= w2_divn1 5 wx wy in (reduce_7 q, reduce_2 r)
- | N7 wx, N3 wy => let (q, r):= w3_divn1 4 wx wy in (reduce_7 q, reduce_3 r)
- | N7 wx, N4 wy => let (q, r):= w4_divn1 3 wx wy in (reduce_7 q, reduce_4 r)
- | N7 wx, N5 wy => let (q, r):= w5_divn1 2 wx wy in (reduce_7 q, reduce_5 r)
- | N7 wx, N6 wy => let (q, r):= w6_divn1 1 wx wy in (reduce_7 q, reduce_6 r)
- | N7 wx, N7 wy => let (q, r):= w7_div_gt wx wy in (reduce_7 q, reduce_7 r)
- | N7 wx, N8 wy =>
- let wx':= GenBase.extend w7_0W 0 wx in
- let (q, r):= w8_div_gt wx' wy in
- (reduce_8 q, reduce_8 r)
- | N7 wx, N9 wy =>
- let wx':= GenBase.extend w7_0W 1 wx in
- let (q, r):= w9_div_gt wx' wy in
- (reduce_9 q, reduce_9 r)
- | N7 wx, N10 wy =>
- let wx':= GenBase.extend w7_0W 2 wx in
- let (q, r):= w10_div_gt wx' wy in
- (reduce_10 q, reduce_10 r)
- | N7 wx, N11 wy =>
- let wx':= GenBase.extend w7_0W 3 wx in
- let (q, r):= w11_div_gt wx' wy in
- (reduce_11 q, reduce_11 r)
- | N7 wx, N12 wy =>
- let wx':= GenBase.extend w7_0W 4 wx in
- let (q, r):= w12_div_gt wx' wy in
- (reduce_12 q, reduce_12 r)
- | N7 wx, N13 wy =>
- let wx':= GenBase.extend w7_0W 5 wx in
- let (q, r):= w13_div_gt wx' wy in
- (reduce_13 q, reduce_13 r)
- | N7 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w7_0W 6 wx) in
- let (q, r):= (make_op n).(znz_div_gt) wx' wy in
- (reduce_n n q, reduce_n n r)
- | N8 wx, N0 wy => let (q, r):= w0_divn1 8 wx wy in (reduce_8 q, reduce_0 r)
- | N8 wx, N1 wy => let (q, r):= w1_divn1 7 wx wy in (reduce_8 q, reduce_1 r)
- | N8 wx, N2 wy => let (q, r):= w2_divn1 6 wx wy in (reduce_8 q, reduce_2 r)
- | N8 wx, N3 wy => let (q, r):= w3_divn1 5 wx wy in (reduce_8 q, reduce_3 r)
- | N8 wx, N4 wy => let (q, r):= w4_divn1 4 wx wy in (reduce_8 q, reduce_4 r)
- | N8 wx, N5 wy => let (q, r):= w5_divn1 3 wx wy in (reduce_8 q, reduce_5 r)
- | N8 wx, N6 wy => let (q, r):= w6_divn1 2 wx wy in (reduce_8 q, reduce_6 r)
- | N8 wx, N7 wy => let (q, r):= w7_divn1 1 wx wy in (reduce_8 q, reduce_7 r)
- | N8 wx, N8 wy => let (q, r):= w8_div_gt wx wy in (reduce_8 q, reduce_8 r)
- | N8 wx, N9 wy =>
- let wx':= GenBase.extend w8_0W 0 wx in
- let (q, r):= w9_div_gt wx' wy in
- (reduce_9 q, reduce_9 r)
- | N8 wx, N10 wy =>
- let wx':= GenBase.extend w8_0W 1 wx in
- let (q, r):= w10_div_gt wx' wy in
- (reduce_10 q, reduce_10 r)
- | N8 wx, N11 wy =>
- let wx':= GenBase.extend w8_0W 2 wx in
- let (q, r):= w11_div_gt wx' wy in
- (reduce_11 q, reduce_11 r)
- | N8 wx, N12 wy =>
- let wx':= GenBase.extend w8_0W 3 wx in
- let (q, r):= w12_div_gt wx' wy in
- (reduce_12 q, reduce_12 r)
- | N8 wx, N13 wy =>
- let wx':= GenBase.extend w8_0W 4 wx in
- let (q, r):= w13_div_gt wx' wy in
- (reduce_13 q, reduce_13 r)
- | N8 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w8_0W 5 wx) in
- let (q, r):= (make_op n).(znz_div_gt) wx' wy in
- (reduce_n n q, reduce_n n r)
- | N9 wx, N0 wy => let (q, r):= w0_divn1 9 wx wy in (reduce_9 q, reduce_0 r)
- | N9 wx, N1 wy => let (q, r):= w1_divn1 8 wx wy in (reduce_9 q, reduce_1 r)
- | N9 wx, N2 wy => let (q, r):= w2_divn1 7 wx wy in (reduce_9 q, reduce_2 r)
- | N9 wx, N3 wy => let (q, r):= w3_divn1 6 wx wy in (reduce_9 q, reduce_3 r)
- | N9 wx, N4 wy => let (q, r):= w4_divn1 5 wx wy in (reduce_9 q, reduce_4 r)
- | N9 wx, N5 wy => let (q, r):= w5_divn1 4 wx wy in (reduce_9 q, reduce_5 r)
- | N9 wx, N6 wy => let (q, r):= w6_divn1 3 wx wy in (reduce_9 q, reduce_6 r)
- | N9 wx, N7 wy => let (q, r):= w7_divn1 2 wx wy in (reduce_9 q, reduce_7 r)
- | N9 wx, N8 wy => let (q, r):= w8_divn1 1 wx wy in (reduce_9 q, reduce_8 r)
- | N9 wx, N9 wy => let (q, r):= w9_div_gt wx wy in (reduce_9 q, reduce_9 r)
- | N9 wx, N10 wy =>
- let wx':= GenBase.extend w9_0W 0 wx in
- let (q, r):= w10_div_gt wx' wy in
- (reduce_10 q, reduce_10 r)
- | N9 wx, N11 wy =>
- let wx':= GenBase.extend w9_0W 1 wx in
- let (q, r):= w11_div_gt wx' wy in
- (reduce_11 q, reduce_11 r)
- | N9 wx, N12 wy =>
- let wx':= GenBase.extend w9_0W 2 wx in
- let (q, r):= w12_div_gt wx' wy in
- (reduce_12 q, reduce_12 r)
- | N9 wx, N13 wy =>
- let wx':= GenBase.extend w9_0W 3 wx in
- let (q, r):= w13_div_gt wx' wy in
- (reduce_13 q, reduce_13 r)
- | N9 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w9_0W 4 wx) in
- let (q, r):= (make_op n).(znz_div_gt) wx' wy in
- (reduce_n n q, reduce_n n r)
- | N10 wx, N0 wy => let (q, r):= w0_divn1 10 wx wy in (reduce_10 q, reduce_0 r)
- | N10 wx, N1 wy => let (q, r):= w1_divn1 9 wx wy in (reduce_10 q, reduce_1 r)
- | N10 wx, N2 wy => let (q, r):= w2_divn1 8 wx wy in (reduce_10 q, reduce_2 r)
- | N10 wx, N3 wy => let (q, r):= w3_divn1 7 wx wy in (reduce_10 q, reduce_3 r)
- | N10 wx, N4 wy => let (q, r):= w4_divn1 6 wx wy in (reduce_10 q, reduce_4 r)
- | N10 wx, N5 wy => let (q, r):= w5_divn1 5 wx wy in (reduce_10 q, reduce_5 r)
- | N10 wx, N6 wy => let (q, r):= w6_divn1 4 wx wy in (reduce_10 q, reduce_6 r)
- | N10 wx, N7 wy => let (q, r):= w7_divn1 3 wx wy in (reduce_10 q, reduce_7 r)
- | N10 wx, N8 wy => let (q, r):= w8_divn1 2 wx wy in (reduce_10 q, reduce_8 r)
- | N10 wx, N9 wy => let (q, r):= w9_divn1 1 wx wy in (reduce_10 q, reduce_9 r)
- | N10 wx, N10 wy => let (q, r):= w10_div_gt wx wy in (reduce_10 q, reduce_10 r)
- | N10 wx, N11 wy =>
- let wx':= GenBase.extend w10_0W 0 wx in
- let (q, r):= w11_div_gt wx' wy in
- (reduce_11 q, reduce_11 r)
- | N10 wx, N12 wy =>
- let wx':= GenBase.extend w10_0W 1 wx in
- let (q, r):= w12_div_gt wx' wy in
- (reduce_12 q, reduce_12 r)
- | N10 wx, N13 wy =>
- let wx':= GenBase.extend w10_0W 2 wx in
- let (q, r):= w13_div_gt wx' wy in
- (reduce_13 q, reduce_13 r)
- | N10 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w10_0W 3 wx) in
- let (q, r):= (make_op n).(znz_div_gt) wx' wy in
- (reduce_n n q, reduce_n n r)
- | N11 wx, N0 wy => let (q, r):= w0_divn1 11 wx wy in (reduce_11 q, reduce_0 r)
- | N11 wx, N1 wy => let (q, r):= w1_divn1 10 wx wy in (reduce_11 q, reduce_1 r)
- | N11 wx, N2 wy => let (q, r):= w2_divn1 9 wx wy in (reduce_11 q, reduce_2 r)
- | N11 wx, N3 wy => let (q, r):= w3_divn1 8 wx wy in (reduce_11 q, reduce_3 r)
- | N11 wx, N4 wy => let (q, r):= w4_divn1 7 wx wy in (reduce_11 q, reduce_4 r)
- | N11 wx, N5 wy => let (q, r):= w5_divn1 6 wx wy in (reduce_11 q, reduce_5 r)
- | N11 wx, N6 wy => let (q, r):= w6_divn1 5 wx wy in (reduce_11 q, reduce_6 r)
- | N11 wx, N7 wy => let (q, r):= w7_divn1 4 wx wy in (reduce_11 q, reduce_7 r)
- | N11 wx, N8 wy => let (q, r):= w8_divn1 3 wx wy in (reduce_11 q, reduce_8 r)
- | N11 wx, N9 wy => let (q, r):= w9_divn1 2 wx wy in (reduce_11 q, reduce_9 r)
- | N11 wx, N10 wy => let (q, r):= w10_divn1 1 wx wy in (reduce_11 q, reduce_10 r)
- | N11 wx, N11 wy => let (q, r):= w11_div_gt wx wy in (reduce_11 q, reduce_11 r)
- | N11 wx, N12 wy =>
- let wx':= GenBase.extend w11_0W 0 wx in
- let (q, r):= w12_div_gt wx' wy in
- (reduce_12 q, reduce_12 r)
- | N11 wx, N13 wy =>
- let wx':= GenBase.extend w11_0W 1 wx in
- let (q, r):= w13_div_gt wx' wy in
- (reduce_13 q, reduce_13 r)
- | N11 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w11_0W 2 wx) in
- let (q, r):= (make_op n).(znz_div_gt) wx' wy in
- (reduce_n n q, reduce_n n r)
- | N12 wx, N0 wy => let (q, r):= w0_divn1 12 wx wy in (reduce_12 q, reduce_0 r)
- | N12 wx, N1 wy => let (q, r):= w1_divn1 11 wx wy in (reduce_12 q, reduce_1 r)
- | N12 wx, N2 wy => let (q, r):= w2_divn1 10 wx wy in (reduce_12 q, reduce_2 r)
- | N12 wx, N3 wy => let (q, r):= w3_divn1 9 wx wy in (reduce_12 q, reduce_3 r)
- | N12 wx, N4 wy => let (q, r):= w4_divn1 8 wx wy in (reduce_12 q, reduce_4 r)
- | N12 wx, N5 wy => let (q, r):= w5_divn1 7 wx wy in (reduce_12 q, reduce_5 r)
- | N12 wx, N6 wy => let (q, r):= w6_divn1 6 wx wy in (reduce_12 q, reduce_6 r)
- | N12 wx, N7 wy => let (q, r):= w7_divn1 5 wx wy in (reduce_12 q, reduce_7 r)
- | N12 wx, N8 wy => let (q, r):= w8_divn1 4 wx wy in (reduce_12 q, reduce_8 r)
- | N12 wx, N9 wy => let (q, r):= w9_divn1 3 wx wy in (reduce_12 q, reduce_9 r)
- | N12 wx, N10 wy => let (q, r):= w10_divn1 2 wx wy in (reduce_12 q, reduce_10 r)
- | N12 wx, N11 wy => let (q, r):= w11_divn1 1 wx wy in (reduce_12 q, reduce_11 r)
- | N12 wx, N12 wy => let (q, r):= w12_div_gt wx wy in (reduce_12 q, reduce_12 r)
- | N12 wx, N13 wy =>
- let wx':= GenBase.extend w12_0W 0 wx in
- let (q, r):= w13_div_gt wx' wy in
- (reduce_13 q, reduce_13 r)
- | N12 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w12_0W 1 wx) in
- let (q, r):= (make_op n).(znz_div_gt) wx' wy in
- (reduce_n n q, reduce_n n r)
- | N13 wx, N0 wy => let (q, r):= w0_divn1 13 wx wy in (reduce_13 q, reduce_0 r)
- | N13 wx, N1 wy => let (q, r):= w1_divn1 12 wx wy in (reduce_13 q, reduce_1 r)
- | N13 wx, N2 wy => let (q, r):= w2_divn1 11 wx wy in (reduce_13 q, reduce_2 r)
- | N13 wx, N3 wy => let (q, r):= w3_divn1 10 wx wy in (reduce_13 q, reduce_3 r)
- | N13 wx, N4 wy => let (q, r):= w4_divn1 9 wx wy in (reduce_13 q, reduce_4 r)
- | N13 wx, N5 wy => let (q, r):= w5_divn1 8 wx wy in (reduce_13 q, reduce_5 r)
- | N13 wx, N6 wy => let (q, r):= w6_divn1 7 wx wy in (reduce_13 q, reduce_6 r)
- | N13 wx, N7 wy => let (q, r):= w7_divn1 6 wx wy in (reduce_13 q, reduce_7 r)
- | N13 wx, N8 wy => let (q, r):= w8_divn1 5 wx wy in (reduce_13 q, reduce_8 r)
- | N13 wx, N9 wy => let (q, r):= w9_divn1 4 wx wy in (reduce_13 q, reduce_9 r)
- | N13 wx, N10 wy => let (q, r):= w10_divn1 3 wx wy in (reduce_13 q, reduce_10 r)
- | N13 wx, N11 wy => let (q, r):= w11_divn1 2 wx wy in (reduce_13 q, reduce_11 r)
- | N13 wx, N12 wy => let (q, r):= w12_divn1 1 wx wy in (reduce_13 q, reduce_12 r)
- | N13 wx, N13 wy => let (q, r):= w13_div_gt wx wy in (reduce_13 q, reduce_13 r)
- | N13 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w13_0W 0 wx) in
- let (q, r):= (make_op n).(znz_div_gt) wx' wy in
- (reduce_n n q, reduce_n n r)
- | Nn n wx, N0 wy =>
- let wy':= GenBase.extend w0_0W 12 wy in
- let (q, r):= w13_divn1 (S n) wx wy' in
- (reduce_n n q, reduce_13 r)
- | Nn n wx, N1 wy =>
- let wy':= GenBase.extend w1_0W 11 wy in
- let (q, r):= w13_divn1 (S n) wx wy' in
- (reduce_n n q, reduce_13 r)
- | Nn n wx, N2 wy =>
- let wy':= GenBase.extend w2_0W 10 wy in
- let (q, r):= w13_divn1 (S n) wx wy' in
- (reduce_n n q, reduce_13 r)
- | Nn n wx, N3 wy =>
- let wy':= GenBase.extend w3_0W 9 wy in
- let (q, r):= w13_divn1 (S n) wx wy' in
- (reduce_n n q, reduce_13 r)
- | Nn n wx, N4 wy =>
- let wy':= GenBase.extend w4_0W 8 wy in
- let (q, r):= w13_divn1 (S n) wx wy' in
- (reduce_n n q, reduce_13 r)
- | Nn n wx, N5 wy =>
- let wy':= GenBase.extend w5_0W 7 wy in
- let (q, r):= w13_divn1 (S n) wx wy' in
- (reduce_n n q, reduce_13 r)
- | Nn n wx, N6 wy =>
- let wy':= GenBase.extend w6_0W 6 wy in
- let (q, r):= w13_divn1 (S n) wx wy' in
- (reduce_n n q, reduce_13 r)
- | Nn n wx, N7 wy =>
- let wy':= GenBase.extend w7_0W 5 wy in
- let (q, r):= w13_divn1 (S n) wx wy' in
- (reduce_n n q, reduce_13 r)
- | Nn n wx, N8 wy =>
- let wy':= GenBase.extend w8_0W 4 wy in
- let (q, r):= w13_divn1 (S n) wx wy' in
- (reduce_n n q, reduce_13 r)
- | Nn n wx, N9 wy =>
- let wy':= GenBase.extend w9_0W 3 wy in
- let (q, r):= w13_divn1 (S n) wx wy' in
- (reduce_n n q, reduce_13 r)
- | Nn n wx, N10 wy =>
- let wy':= GenBase.extend w10_0W 2 wy in
- let (q, r):= w13_divn1 (S n) wx wy' in
- (reduce_n n q, reduce_13 r)
- | Nn n wx, N11 wy =>
- let wy':= GenBase.extend w11_0W 1 wy in
- let (q, r):= w13_divn1 (S n) wx wy' in
- (reduce_n n q, reduce_13 r)
- | Nn n wx, N12 wy =>
- let wy':= GenBase.extend w12_0W 0 wy in
- let (q, r):= w13_divn1 (S n) wx wy' in
- (reduce_n n q, reduce_13 r)
- | Nn n wx, N13 wy =>
- let wy':= wy in
- let (q, r):= w13_divn1 (S n) wx wy' in
- (reduce_n n q, reduce_13 r)
- | Nn n wx, Nn m wy =>
+ w13_op.(znz_compare) w13_op.(znz_sub) (S n) x y in
+ (Nn _ u, N13 v).
+
+ Let div_gt0 x y := let (u,v) := (w0_div_gt x y) in (reduce_0 u, reduce_0 v).
+ Let div_gt1 x y := let (u,v) := (w1_div_gt x y) in (reduce_1 u, reduce_1 v).
+ Let div_gt2 x y := let (u,v) := (w2_div_gt x y) in (reduce_2 u, reduce_2 v).
+ Let div_gt3 x y := let (u,v) := (w3_div_gt x y) in (reduce_3 u, reduce_3 v).
+ Let div_gt4 x y := let (u,v) := (w4_div_gt x y) in (reduce_4 u, reduce_4 v).
+ Let div_gt5 x y := let (u,v) := (w5_div_gt x y) in (reduce_5 u, reduce_5 v).
+ Let div_gt6 x y := let (u,v) := (w6_div_gt x y) in (reduce_6 u, reduce_6 v).
+ Let div_gt7 x y := let (u,v) := (w7_div_gt x y) in (reduce_7 u, reduce_7 v).
+ Let div_gt8 x y := let (u,v) := (w8_div_gt x y) in (reduce_8 u, reduce_8 v).
+ Let div_gt9 x y := let (u,v) := (w9_div_gt x y) in (reduce_9 u, reduce_9 v).
+ Let div_gt10 x y := let (u,v) := (w10_div_gt x y) in (reduce_10 u, reduce_10 v).
+ Let div_gt11 x y := let (u,v) := (w11_div_gt x y) in (reduce_11 u, reduce_11 v).
+ Let div_gt12 x y := let (u,v) := (w12_div_gt x y) in (reduce_12 u, reduce_12 v).
+ Let div_gt13 x y := let (u,v) := (w13_div_gt x y) in (reduce_13 u, reduce_13 v).
+
+ Let div_gtnm n m wx wy :=
let mn := Max.max n m in
let d := diff n m in
let op := make_op mn in
- let (q, r):= op.(znz_div)
+ let (q, r):= op.(znz_div_gt)
(castm (diff_r n m) (extend_tr wx (snd d)))
(castm (diff_l n m) (extend_tr wy (fst d))) in
- (reduce_n mn q, reduce_n mn r)
- end.
+ (reduce_n mn q, reduce_n mn r).
+
+ Definition div_gt := Eval lazy beta delta [iter] in
+ (iter _
+ div_gt0
+ (fun n x y => div_gt0 x (GenBase.get_low w_0 (S n) y))
+ w0_divn1
+ div_gt1
+ (fun n x y => div_gt1 x (GenBase.get_low W0 (S n) y))
+ w1_divn1
+ div_gt2
+ (fun n x y => div_gt2 x (GenBase.get_low W0 (S n) y))
+ w2_divn1
+ div_gt3
+ (fun n x y => div_gt3 x (GenBase.get_low W0 (S n) y))
+ w3_divn1
+ div_gt4
+ (fun n x y => div_gt4 x (GenBase.get_low W0 (S n) y))
+ w4_divn1
+ div_gt5
+ (fun n x y => div_gt5 x (GenBase.get_low W0 (S n) y))
+ w5_divn1
+ div_gt6
+ (fun n x y => div_gt6 x (GenBase.get_low W0 (S n) y))
+ w6_divn1
+ div_gt7
+ (fun n x y => div_gt7 x (GenBase.get_low W0 (S n) y))
+ w7_divn1
+ div_gt8
+ (fun n x y => div_gt8 x (GenBase.get_low W0 (S n) y))
+ w8_divn1
+ div_gt9
+ (fun n x y => div_gt9 x (GenBase.get_low W0 (S n) y))
+ w9_divn1
+ div_gt10
+ (fun n x y => div_gt10 x (GenBase.get_low W0 (S n) y))
+ w10_divn1
+ div_gt11
+ (fun n x y => div_gt11 x (GenBase.get_low W0 (S n) y))
+ w11_divn1
+ div_gt12
+ (fun n x y => div_gt12 x (GenBase.get_low W0 (S n) y))
+ w12_divn1
+ div_gt13
+ (fun n x y => div_gt13 x (GenBase.get_low W0 (S n) y))
+ w13_divn1
+ div_gtnm).
+
+ Theorem spec_div_gt: forall x y,
+ [x] > [y] -> 0 < [y] ->
+ let (q,r) := div_gt x y in
+ [q] = [x] / [y] /\ [r] = [x] mod [y].
+ Admitted.
Definition div_eucl x y :=
match compare x y with
@@ -3244,8 +2909,24 @@ Module Make (W0:W0Type).
| Gt => div_gt x y
end.
+ Theorem spec_div_eucl: forall x y,
+ 0 < [y] ->
+ let (q,r) := div_eucl x y in
+ [q] = [x] / [y] /\ [r] = [x] mod [y].
+ Admitted.
+
Definition div x y := fst (div_eucl x y).
+ Theorem spec_div:
+ forall x y, 0 < [y] -> [div x y] = [x] / [y].
+ Admitted.
+
+ (***************************************************************)
+ (* *)
+ (* Modulo *)
+ (* *)
+ (***************************************************************)
+
Definition w0_mod_gt := w0_op.(znz_mod_gt).
Definition w1_mod_gt := w1_op.(znz_mod_gt).
Definition w2_mod_gt := w2_op.(znz_mod_gt).
@@ -3318,478 +2999,63 @@ Module Make (W0:W0Type).
w13_op.(znz_head0) w13_op.(znz_add_mul_div) w13_op.(znz_div21)
w13_op.(znz_compare) w13_op.(znz_sub).
- Definition mod_gt x y :=
- match x, y with
- | N0 wx, N0 wy => reduce_0 (w0_mod_gt wx wy)
- | N0 wx, N1 wy =>
- let wx':= GenBase.extend w0_0W 0 wx in
- reduce_1 (w1_mod_gt wx' wy)
- | N0 wx, N2 wy =>
- let wx':= GenBase.extend w0_0W 1 wx in
- reduce_2 (w2_mod_gt wx' wy)
- | N0 wx, N3 wy =>
- let wx':= GenBase.extend w0_0W 2 wx in
- reduce_3 (w3_mod_gt wx' wy)
- | N0 wx, N4 wy =>
- let wx':= GenBase.extend w0_0W 3 wx in
- reduce_4 (w4_mod_gt wx' wy)
- | N0 wx, N5 wy =>
- let wx':= GenBase.extend w0_0W 4 wx in
- reduce_5 (w5_mod_gt wx' wy)
- | N0 wx, N6 wy =>
- let wx':= GenBase.extend w0_0W 5 wx in
- reduce_6 (w6_mod_gt wx' wy)
- | N0 wx, N7 wy =>
- let wx':= GenBase.extend w0_0W 6 wx in
- reduce_7 (w7_mod_gt wx' wy)
- | N0 wx, N8 wy =>
- let wx':= GenBase.extend w0_0W 7 wx in
- reduce_8 (w8_mod_gt wx' wy)
- | N0 wx, N9 wy =>
- let wx':= GenBase.extend w0_0W 8 wx in
- reduce_9 (w9_mod_gt wx' wy)
- | N0 wx, N10 wy =>
- let wx':= GenBase.extend w0_0W 9 wx in
- reduce_10 (w10_mod_gt wx' wy)
- | N0 wx, N11 wy =>
- let wx':= GenBase.extend w0_0W 10 wx in
- reduce_11 (w11_mod_gt wx' wy)
- | N0 wx, N12 wy =>
- let wx':= GenBase.extend w0_0W 11 wx in
- reduce_12 (w12_mod_gt wx' wy)
- | N0 wx, N13 wy =>
- let wx':= GenBase.extend w0_0W 12 wx in
- reduce_13 (w13_mod_gt wx' wy)
- | N0 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w0_0W 13 wx) in
- reduce_n n ((make_op n).(znz_mod_gt) wx' wy)
- | N1 wx, N0 wy => reduce_0 (w0_modn1 1 wx wy)
- | N1 wx, N1 wy => reduce_1 (w1_mod_gt wx wy)
- | N1 wx, N2 wy =>
- let wx':= GenBase.extend w1_0W 0 wx in
- reduce_2 (w2_mod_gt wx' wy)
- | N1 wx, N3 wy =>
- let wx':= GenBase.extend w1_0W 1 wx in
- reduce_3 (w3_mod_gt wx' wy)
- | N1 wx, N4 wy =>
- let wx':= GenBase.extend w1_0W 2 wx in
- reduce_4 (w4_mod_gt wx' wy)
- | N1 wx, N5 wy =>
- let wx':= GenBase.extend w1_0W 3 wx in
- reduce_5 (w5_mod_gt wx' wy)
- | N1 wx, N6 wy =>
- let wx':= GenBase.extend w1_0W 4 wx in
- reduce_6 (w6_mod_gt wx' wy)
- | N1 wx, N7 wy =>
- let wx':= GenBase.extend w1_0W 5 wx in
- reduce_7 (w7_mod_gt wx' wy)
- | N1 wx, N8 wy =>
- let wx':= GenBase.extend w1_0W 6 wx in
- reduce_8 (w8_mod_gt wx' wy)
- | N1 wx, N9 wy =>
- let wx':= GenBase.extend w1_0W 7 wx in
- reduce_9 (w9_mod_gt wx' wy)
- | N1 wx, N10 wy =>
- let wx':= GenBase.extend w1_0W 8 wx in
- reduce_10 (w10_mod_gt wx' wy)
- | N1 wx, N11 wy =>
- let wx':= GenBase.extend w1_0W 9 wx in
- reduce_11 (w11_mod_gt wx' wy)
- | N1 wx, N12 wy =>
- let wx':= GenBase.extend w1_0W 10 wx in
- reduce_12 (w12_mod_gt wx' wy)
- | N1 wx, N13 wy =>
- let wx':= GenBase.extend w1_0W 11 wx in
- reduce_13 (w13_mod_gt wx' wy)
- | N1 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w1_0W 12 wx) in
- reduce_n n ((make_op n).(znz_mod_gt) wx' wy)
- | N2 wx, N0 wy => reduce_0 (w0_modn1 2 wx wy)
- | N2 wx, N1 wy => reduce_1 (w1_modn1 1 wx wy)
- | N2 wx, N2 wy => reduce_2 (w2_mod_gt wx wy)
- | N2 wx, N3 wy =>
- let wx':= GenBase.extend w2_0W 0 wx in
- reduce_3 (w3_mod_gt wx' wy)
- | N2 wx, N4 wy =>
- let wx':= GenBase.extend w2_0W 1 wx in
- reduce_4 (w4_mod_gt wx' wy)
- | N2 wx, N5 wy =>
- let wx':= GenBase.extend w2_0W 2 wx in
- reduce_5 (w5_mod_gt wx' wy)
- | N2 wx, N6 wy =>
- let wx':= GenBase.extend w2_0W 3 wx in
- reduce_6 (w6_mod_gt wx' wy)
- | N2 wx, N7 wy =>
- let wx':= GenBase.extend w2_0W 4 wx in
- reduce_7 (w7_mod_gt wx' wy)
- | N2 wx, N8 wy =>
- let wx':= GenBase.extend w2_0W 5 wx in
- reduce_8 (w8_mod_gt wx' wy)
- | N2 wx, N9 wy =>
- let wx':= GenBase.extend w2_0W 6 wx in
- reduce_9 (w9_mod_gt wx' wy)
- | N2 wx, N10 wy =>
- let wx':= GenBase.extend w2_0W 7 wx in
- reduce_10 (w10_mod_gt wx' wy)
- | N2 wx, N11 wy =>
- let wx':= GenBase.extend w2_0W 8 wx in
- reduce_11 (w11_mod_gt wx' wy)
- | N2 wx, N12 wy =>
- let wx':= GenBase.extend w2_0W 9 wx in
- reduce_12 (w12_mod_gt wx' wy)
- | N2 wx, N13 wy =>
- let wx':= GenBase.extend w2_0W 10 wx in
- reduce_13 (w13_mod_gt wx' wy)
- | N2 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w2_0W 11 wx) in
- reduce_n n ((make_op n).(znz_mod_gt) wx' wy)
- | N3 wx, N0 wy => reduce_0 (w0_modn1 3 wx wy)
- | N3 wx, N1 wy => reduce_1 (w1_modn1 2 wx wy)
- | N3 wx, N2 wy => reduce_2 (w2_modn1 1 wx wy)
- | N3 wx, N3 wy => reduce_3 (w3_mod_gt wx wy)
- | N3 wx, N4 wy =>
- let wx':= GenBase.extend w3_0W 0 wx in
- reduce_4 (w4_mod_gt wx' wy)
- | N3 wx, N5 wy =>
- let wx':= GenBase.extend w3_0W 1 wx in
- reduce_5 (w5_mod_gt wx' wy)
- | N3 wx, N6 wy =>
- let wx':= GenBase.extend w3_0W 2 wx in
- reduce_6 (w6_mod_gt wx' wy)
- | N3 wx, N7 wy =>
- let wx':= GenBase.extend w3_0W 3 wx in
- reduce_7 (w7_mod_gt wx' wy)
- | N3 wx, N8 wy =>
- let wx':= GenBase.extend w3_0W 4 wx in
- reduce_8 (w8_mod_gt wx' wy)
- | N3 wx, N9 wy =>
- let wx':= GenBase.extend w3_0W 5 wx in
- reduce_9 (w9_mod_gt wx' wy)
- | N3 wx, N10 wy =>
- let wx':= GenBase.extend w3_0W 6 wx in
- reduce_10 (w10_mod_gt wx' wy)
- | N3 wx, N11 wy =>
- let wx':= GenBase.extend w3_0W 7 wx in
- reduce_11 (w11_mod_gt wx' wy)
- | N3 wx, N12 wy =>
- let wx':= GenBase.extend w3_0W 8 wx in
- reduce_12 (w12_mod_gt wx' wy)
- | N3 wx, N13 wy =>
- let wx':= GenBase.extend w3_0W 9 wx in
- reduce_13 (w13_mod_gt wx' wy)
- | N3 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w3_0W 10 wx) in
- reduce_n n ((make_op n).(znz_mod_gt) wx' wy)
- | N4 wx, N0 wy => reduce_0 (w0_modn1 4 wx wy)
- | N4 wx, N1 wy => reduce_1 (w1_modn1 3 wx wy)
- | N4 wx, N2 wy => reduce_2 (w2_modn1 2 wx wy)
- | N4 wx, N3 wy => reduce_3 (w3_modn1 1 wx wy)
- | N4 wx, N4 wy => reduce_4 (w4_mod_gt wx wy)
- | N4 wx, N5 wy =>
- let wx':= GenBase.extend w4_0W 0 wx in
- reduce_5 (w5_mod_gt wx' wy)
- | N4 wx, N6 wy =>
- let wx':= GenBase.extend w4_0W 1 wx in
- reduce_6 (w6_mod_gt wx' wy)
- | N4 wx, N7 wy =>
- let wx':= GenBase.extend w4_0W 2 wx in
- reduce_7 (w7_mod_gt wx' wy)
- | N4 wx, N8 wy =>
- let wx':= GenBase.extend w4_0W 3 wx in
- reduce_8 (w8_mod_gt wx' wy)
- | N4 wx, N9 wy =>
- let wx':= GenBase.extend w4_0W 4 wx in
- reduce_9 (w9_mod_gt wx' wy)
- | N4 wx, N10 wy =>
- let wx':= GenBase.extend w4_0W 5 wx in
- reduce_10 (w10_mod_gt wx' wy)
- | N4 wx, N11 wy =>
- let wx':= GenBase.extend w4_0W 6 wx in
- reduce_11 (w11_mod_gt wx' wy)
- | N4 wx, N12 wy =>
- let wx':= GenBase.extend w4_0W 7 wx in
- reduce_12 (w12_mod_gt wx' wy)
- | N4 wx, N13 wy =>
- let wx':= GenBase.extend w4_0W 8 wx in
- reduce_13 (w13_mod_gt wx' wy)
- | N4 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w4_0W 9 wx) in
- reduce_n n ((make_op n).(znz_mod_gt) wx' wy)
- | N5 wx, N0 wy => reduce_0 (w0_modn1 5 wx wy)
- | N5 wx, N1 wy => reduce_1 (w1_modn1 4 wx wy)
- | N5 wx, N2 wy => reduce_2 (w2_modn1 3 wx wy)
- | N5 wx, N3 wy => reduce_3 (w3_modn1 2 wx wy)
- | N5 wx, N4 wy => reduce_4 (w4_modn1 1 wx wy)
- | N5 wx, N5 wy => reduce_5 (w5_mod_gt wx wy)
- | N5 wx, N6 wy =>
- let wx':= GenBase.extend w5_0W 0 wx in
- reduce_6 (w6_mod_gt wx' wy)
- | N5 wx, N7 wy =>
- let wx':= GenBase.extend w5_0W 1 wx in
- reduce_7 (w7_mod_gt wx' wy)
- | N5 wx, N8 wy =>
- let wx':= GenBase.extend w5_0W 2 wx in
- reduce_8 (w8_mod_gt wx' wy)
- | N5 wx, N9 wy =>
- let wx':= GenBase.extend w5_0W 3 wx in
- reduce_9 (w9_mod_gt wx' wy)
- | N5 wx, N10 wy =>
- let wx':= GenBase.extend w5_0W 4 wx in
- reduce_10 (w10_mod_gt wx' wy)
- | N5 wx, N11 wy =>
- let wx':= GenBase.extend w5_0W 5 wx in
- reduce_11 (w11_mod_gt wx' wy)
- | N5 wx, N12 wy =>
- let wx':= GenBase.extend w5_0W 6 wx in
- reduce_12 (w12_mod_gt wx' wy)
- | N5 wx, N13 wy =>
- let wx':= GenBase.extend w5_0W 7 wx in
- reduce_13 (w13_mod_gt wx' wy)
- | N5 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w5_0W 8 wx) in
- reduce_n n ((make_op n).(znz_mod_gt) wx' wy)
- | N6 wx, N0 wy => reduce_0 (w0_modn1 6 wx wy)
- | N6 wx, N1 wy => reduce_1 (w1_modn1 5 wx wy)
- | N6 wx, N2 wy => reduce_2 (w2_modn1 4 wx wy)
- | N6 wx, N3 wy => reduce_3 (w3_modn1 3 wx wy)
- | N6 wx, N4 wy => reduce_4 (w4_modn1 2 wx wy)
- | N6 wx, N5 wy => reduce_5 (w5_modn1 1 wx wy)
- | N6 wx, N6 wy => reduce_6 (w6_mod_gt wx wy)
- | N6 wx, N7 wy =>
- let wx':= GenBase.extend w6_0W 0 wx in
- reduce_7 (w7_mod_gt wx' wy)
- | N6 wx, N8 wy =>
- let wx':= GenBase.extend w6_0W 1 wx in
- reduce_8 (w8_mod_gt wx' wy)
- | N6 wx, N9 wy =>
- let wx':= GenBase.extend w6_0W 2 wx in
- reduce_9 (w9_mod_gt wx' wy)
- | N6 wx, N10 wy =>
- let wx':= GenBase.extend w6_0W 3 wx in
- reduce_10 (w10_mod_gt wx' wy)
- | N6 wx, N11 wy =>
- let wx':= GenBase.extend w6_0W 4 wx in
- reduce_11 (w11_mod_gt wx' wy)
- | N6 wx, N12 wy =>
- let wx':= GenBase.extend w6_0W 5 wx in
- reduce_12 (w12_mod_gt wx' wy)
- | N6 wx, N13 wy =>
- let wx':= GenBase.extend w6_0W 6 wx in
- reduce_13 (w13_mod_gt wx' wy)
- | N6 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w6_0W 7 wx) in
- reduce_n n ((make_op n).(znz_mod_gt) wx' wy)
- | N7 wx, N0 wy => reduce_0 (w0_modn1 7 wx wy)
- | N7 wx, N1 wy => reduce_1 (w1_modn1 6 wx wy)
- | N7 wx, N2 wy => reduce_2 (w2_modn1 5 wx wy)
- | N7 wx, N3 wy => reduce_3 (w3_modn1 4 wx wy)
- | N7 wx, N4 wy => reduce_4 (w4_modn1 3 wx wy)
- | N7 wx, N5 wy => reduce_5 (w5_modn1 2 wx wy)
- | N7 wx, N6 wy => reduce_6 (w6_modn1 1 wx wy)
- | N7 wx, N7 wy => reduce_7 (w7_mod_gt wx wy)
- | N7 wx, N8 wy =>
- let wx':= GenBase.extend w7_0W 0 wx in
- reduce_8 (w8_mod_gt wx' wy)
- | N7 wx, N9 wy =>
- let wx':= GenBase.extend w7_0W 1 wx in
- reduce_9 (w9_mod_gt wx' wy)
- | N7 wx, N10 wy =>
- let wx':= GenBase.extend w7_0W 2 wx in
- reduce_10 (w10_mod_gt wx' wy)
- | N7 wx, N11 wy =>
- let wx':= GenBase.extend w7_0W 3 wx in
- reduce_11 (w11_mod_gt wx' wy)
- | N7 wx, N12 wy =>
- let wx':= GenBase.extend w7_0W 4 wx in
- reduce_12 (w12_mod_gt wx' wy)
- | N7 wx, N13 wy =>
- let wx':= GenBase.extend w7_0W 5 wx in
- reduce_13 (w13_mod_gt wx' wy)
- | N7 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w7_0W 6 wx) in
- reduce_n n ((make_op n).(znz_mod_gt) wx' wy)
- | N8 wx, N0 wy => reduce_0 (w0_modn1 8 wx wy)
- | N8 wx, N1 wy => reduce_1 (w1_modn1 7 wx wy)
- | N8 wx, N2 wy => reduce_2 (w2_modn1 6 wx wy)
- | N8 wx, N3 wy => reduce_3 (w3_modn1 5 wx wy)
- | N8 wx, N4 wy => reduce_4 (w4_modn1 4 wx wy)
- | N8 wx, N5 wy => reduce_5 (w5_modn1 3 wx wy)
- | N8 wx, N6 wy => reduce_6 (w6_modn1 2 wx wy)
- | N8 wx, N7 wy => reduce_7 (w7_modn1 1 wx wy)
- | N8 wx, N8 wy => reduce_8 (w8_mod_gt wx wy)
- | N8 wx, N9 wy =>
- let wx':= GenBase.extend w8_0W 0 wx in
- reduce_9 (w9_mod_gt wx' wy)
- | N8 wx, N10 wy =>
- let wx':= GenBase.extend w8_0W 1 wx in
- reduce_10 (w10_mod_gt wx' wy)
- | N8 wx, N11 wy =>
- let wx':= GenBase.extend w8_0W 2 wx in
- reduce_11 (w11_mod_gt wx' wy)
- | N8 wx, N12 wy =>
- let wx':= GenBase.extend w8_0W 3 wx in
- reduce_12 (w12_mod_gt wx' wy)
- | N8 wx, N13 wy =>
- let wx':= GenBase.extend w8_0W 4 wx in
- reduce_13 (w13_mod_gt wx' wy)
- | N8 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w8_0W 5 wx) in
- reduce_n n ((make_op n).(znz_mod_gt) wx' wy)
- | N9 wx, N0 wy => reduce_0 (w0_modn1 9 wx wy)
- | N9 wx, N1 wy => reduce_1 (w1_modn1 8 wx wy)
- | N9 wx, N2 wy => reduce_2 (w2_modn1 7 wx wy)
- | N9 wx, N3 wy => reduce_3 (w3_modn1 6 wx wy)
- | N9 wx, N4 wy => reduce_4 (w4_modn1 5 wx wy)
- | N9 wx, N5 wy => reduce_5 (w5_modn1 4 wx wy)
- | N9 wx, N6 wy => reduce_6 (w6_modn1 3 wx wy)
- | N9 wx, N7 wy => reduce_7 (w7_modn1 2 wx wy)
- | N9 wx, N8 wy => reduce_8 (w8_modn1 1 wx wy)
- | N9 wx, N9 wy => reduce_9 (w9_mod_gt wx wy)
- | N9 wx, N10 wy =>
- let wx':= GenBase.extend w9_0W 0 wx in
- reduce_10 (w10_mod_gt wx' wy)
- | N9 wx, N11 wy =>
- let wx':= GenBase.extend w9_0W 1 wx in
- reduce_11 (w11_mod_gt wx' wy)
- | N9 wx, N12 wy =>
- let wx':= GenBase.extend w9_0W 2 wx in
- reduce_12 (w12_mod_gt wx' wy)
- | N9 wx, N13 wy =>
- let wx':= GenBase.extend w9_0W 3 wx in
- reduce_13 (w13_mod_gt wx' wy)
- | N9 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w9_0W 4 wx) in
- reduce_n n ((make_op n).(znz_mod_gt) wx' wy)
- | N10 wx, N0 wy => reduce_0 (w0_modn1 10 wx wy)
- | N10 wx, N1 wy => reduce_1 (w1_modn1 9 wx wy)
- | N10 wx, N2 wy => reduce_2 (w2_modn1 8 wx wy)
- | N10 wx, N3 wy => reduce_3 (w3_modn1 7 wx wy)
- | N10 wx, N4 wy => reduce_4 (w4_modn1 6 wx wy)
- | N10 wx, N5 wy => reduce_5 (w5_modn1 5 wx wy)
- | N10 wx, N6 wy => reduce_6 (w6_modn1 4 wx wy)
- | N10 wx, N7 wy => reduce_7 (w7_modn1 3 wx wy)
- | N10 wx, N8 wy => reduce_8 (w8_modn1 2 wx wy)
- | N10 wx, N9 wy => reduce_9 (w9_modn1 1 wx wy)
- | N10 wx, N10 wy => reduce_10 (w10_mod_gt wx wy)
- | N10 wx, N11 wy =>
- let wx':= GenBase.extend w10_0W 0 wx in
- reduce_11 (w11_mod_gt wx' wy)
- | N10 wx, N12 wy =>
- let wx':= GenBase.extend w10_0W 1 wx in
- reduce_12 (w12_mod_gt wx' wy)
- | N10 wx, N13 wy =>
- let wx':= GenBase.extend w10_0W 2 wx in
- reduce_13 (w13_mod_gt wx' wy)
- | N10 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w10_0W 3 wx) in
- reduce_n n ((make_op n).(znz_mod_gt) wx' wy)
- | N11 wx, N0 wy => reduce_0 (w0_modn1 11 wx wy)
- | N11 wx, N1 wy => reduce_1 (w1_modn1 10 wx wy)
- | N11 wx, N2 wy => reduce_2 (w2_modn1 9 wx wy)
- | N11 wx, N3 wy => reduce_3 (w3_modn1 8 wx wy)
- | N11 wx, N4 wy => reduce_4 (w4_modn1 7 wx wy)
- | N11 wx, N5 wy => reduce_5 (w5_modn1 6 wx wy)
- | N11 wx, N6 wy => reduce_6 (w6_modn1 5 wx wy)
- | N11 wx, N7 wy => reduce_7 (w7_modn1 4 wx wy)
- | N11 wx, N8 wy => reduce_8 (w8_modn1 3 wx wy)
- | N11 wx, N9 wy => reduce_9 (w9_modn1 2 wx wy)
- | N11 wx, N10 wy => reduce_10 (w10_modn1 1 wx wy)
- | N11 wx, N11 wy => reduce_11 (w11_mod_gt wx wy)
- | N11 wx, N12 wy =>
- let wx':= GenBase.extend w11_0W 0 wx in
- reduce_12 (w12_mod_gt wx' wy)
- | N11 wx, N13 wy =>
- let wx':= GenBase.extend w11_0W 1 wx in
- reduce_13 (w13_mod_gt wx' wy)
- | N11 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w11_0W 2 wx) in
- reduce_n n ((make_op n).(znz_mod_gt) wx' wy)
- | N12 wx, N0 wy => reduce_0 (w0_modn1 12 wx wy)
- | N12 wx, N1 wy => reduce_1 (w1_modn1 11 wx wy)
- | N12 wx, N2 wy => reduce_2 (w2_modn1 10 wx wy)
- | N12 wx, N3 wy => reduce_3 (w3_modn1 9 wx wy)
- | N12 wx, N4 wy => reduce_4 (w4_modn1 8 wx wy)
- | N12 wx, N5 wy => reduce_5 (w5_modn1 7 wx wy)
- | N12 wx, N6 wy => reduce_6 (w6_modn1 6 wx wy)
- | N12 wx, N7 wy => reduce_7 (w7_modn1 5 wx wy)
- | N12 wx, N8 wy => reduce_8 (w8_modn1 4 wx wy)
- | N12 wx, N9 wy => reduce_9 (w9_modn1 3 wx wy)
- | N12 wx, N10 wy => reduce_10 (w10_modn1 2 wx wy)
- | N12 wx, N11 wy => reduce_11 (w11_modn1 1 wx wy)
- | N12 wx, N12 wy => reduce_12 (w12_mod_gt wx wy)
- | N12 wx, N13 wy =>
- let wx':= GenBase.extend w12_0W 0 wx in
- reduce_13 (w13_mod_gt wx' wy)
- | N12 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w12_0W 1 wx) in
- reduce_n n ((make_op n).(znz_mod_gt) wx' wy)
- | N13 wx, N0 wy => reduce_0 (w0_modn1 13 wx wy)
- | N13 wx, N1 wy => reduce_1 (w1_modn1 12 wx wy)
- | N13 wx, N2 wy => reduce_2 (w2_modn1 11 wx wy)
- | N13 wx, N3 wy => reduce_3 (w3_modn1 10 wx wy)
- | N13 wx, N4 wy => reduce_4 (w4_modn1 9 wx wy)
- | N13 wx, N5 wy => reduce_5 (w5_modn1 8 wx wy)
- | N13 wx, N6 wy => reduce_6 (w6_modn1 7 wx wy)
- | N13 wx, N7 wy => reduce_7 (w7_modn1 6 wx wy)
- | N13 wx, N8 wy => reduce_8 (w8_modn1 5 wx wy)
- | N13 wx, N9 wy => reduce_9 (w9_modn1 4 wx wy)
- | N13 wx, N10 wy => reduce_10 (w10_modn1 3 wx wy)
- | N13 wx, N11 wy => reduce_11 (w11_modn1 2 wx wy)
- | N13 wx, N12 wy => reduce_12 (w12_modn1 1 wx wy)
- | N13 wx, N13 wy => reduce_13 (w13_mod_gt wx wy)
- | N13 wx, Nn n wy =>
- let wx':= extend n w13 (GenBase.extend w13_0W 0 wx) in
- reduce_n n ((make_op n).(znz_mod_gt) wx' wy)
- | Nn n wx, N0 wy =>
- let wy':= GenBase.extend w0_0W 12 wy in
- reduce_13 (w13_modn1 (S n) wx wy')
- | Nn n wx, N1 wy =>
- let wy':= GenBase.extend w1_0W 11 wy in
- reduce_13 (w13_modn1 (S n) wx wy')
- | Nn n wx, N2 wy =>
- let wy':= GenBase.extend w2_0W 10 wy in
- reduce_13 (w13_modn1 (S n) wx wy')
- | Nn n wx, N3 wy =>
- let wy':= GenBase.extend w3_0W 9 wy in
- reduce_13 (w13_modn1 (S n) wx wy')
- | Nn n wx, N4 wy =>
- let wy':= GenBase.extend w4_0W 8 wy in
- reduce_13 (w13_modn1 (S n) wx wy')
- | Nn n wx, N5 wy =>
- let wy':= GenBase.extend w5_0W 7 wy in
- reduce_13 (w13_modn1 (S n) wx wy')
- | Nn n wx, N6 wy =>
- let wy':= GenBase.extend w6_0W 6 wy in
- reduce_13 (w13_modn1 (S n) wx wy')
- | Nn n wx, N7 wy =>
- let wy':= GenBase.extend w7_0W 5 wy in
- reduce_13 (w13_modn1 (S n) wx wy')
- | Nn n wx, N8 wy =>
- let wy':= GenBase.extend w8_0W 4 wy in
- reduce_13 (w13_modn1 (S n) wx wy')
- | Nn n wx, N9 wy =>
- let wy':= GenBase.extend w9_0W 3 wy in
- reduce_13 (w13_modn1 (S n) wx wy')
- | Nn n wx, N10 wy =>
- let wy':= GenBase.extend w10_0W 2 wy in
- reduce_13 (w13_modn1 (S n) wx wy')
- | Nn n wx, N11 wy =>
- let wy':= GenBase.extend w11_0W 1 wy in
- reduce_13 (w13_modn1 (S n) wx wy')
- | Nn n wx, N12 wy =>
- let wy':= GenBase.extend w12_0W 0 wy in
- reduce_13 (w13_modn1 (S n) wx wy')
- | Nn n wx, N13 wy =>
- let wy':= wy in
- reduce_13 (w13_modn1 (S n) wx wy')
- | Nn n wx, Nn m wy =>
+ Let mod_gtnm n m wx wy :=
let mn := Max.max n m in
let d := diff n m in
let op := make_op mn in
- reduce_n mn (op.(znz_mod_gt)
- (castm (diff_r n m) (extend_tr wx (snd d)))
- (castm (diff_l n m) (extend_tr wy (fst d))))
- end.
+ reduce_n mn (op.(znz_mod_gt)
+ (castm (diff_r n m) (extend_tr wx (snd d)))
+ (castm (diff_l n m) (extend_tr wy (fst d)))).
+
+ Definition mod_gt := Eval lazy beta delta[iter] in
+ (iter _
+ (fun x y => reduce_0 (w0_mod_gt x y))
+ (fun n x y => reduce_0 (w0_mod_gt x (GenBase.get_low w_0 (S n) y)))
+ (fun n x y => reduce_0 (w0_modn1 (S n) x y))
+ (fun x y => reduce_1 (w1_mod_gt x y))
+ (fun n x y => reduce_1 (w1_mod_gt x (GenBase.get_low W0 (S n) y)))
+ (fun n x y => reduce_1 (w1_modn1 (S n) x y))
+ (fun x y => reduce_2 (w2_mod_gt x y))
+ (fun n x y => reduce_2 (w2_mod_gt x (GenBase.get_low W0 (S n) y)))
+ (fun n x y => reduce_2 (w2_modn1 (S n) x y))
+ (fun x y => reduce_3 (w3_mod_gt x y))
+ (fun n x y => reduce_3 (w3_mod_gt x (GenBase.get_low W0 (S n) y)))
+ (fun n x y => reduce_3 (w3_modn1 (S n) x y))
+ (fun x y => reduce_4 (w4_mod_gt x y))
+ (fun n x y => reduce_4 (w4_mod_gt x (GenBase.get_low W0 (S n) y)))
+ (fun n x y => reduce_4 (w4_modn1 (S n) x y))
+ (fun x y => reduce_5 (w5_mod_gt x y))
+ (fun n x y => reduce_5 (w5_mod_gt x (GenBase.get_low W0 (S n) y)))
+ (fun n x y => reduce_5 (w5_modn1 (S n) x y))
+ (fun x y => reduce_6 (w6_mod_gt x y))
+ (fun n x y => reduce_6 (w6_mod_gt x (GenBase.get_low W0 (S n) y)))
+ (fun n x y => reduce_6 (w6_modn1 (S n) x y))
+ (fun x y => reduce_7 (w7_mod_gt x y))
+ (fun n x y => reduce_7 (w7_mod_gt x (GenBase.get_low W0 (S n) y)))
+ (fun n x y => reduce_7 (w7_modn1 (S n) x y))
+ (fun x y => reduce_8 (w8_mod_gt x y))
+ (fun n x y => reduce_8 (w8_mod_gt x (GenBase.get_low W0 (S n) y)))
+ (fun n x y => reduce_8 (w8_modn1 (S n) x y))
+ (fun x y => reduce_9 (w9_mod_gt x y))
+ (fun n x y => reduce_9 (w9_mod_gt x (GenBase.get_low W0 (S n) y)))
+ (fun n x y => reduce_9 (w9_modn1 (S n) x y))
+ (fun x y => reduce_10 (w10_mod_gt x y))
+ (fun n x y => reduce_10 (w10_mod_gt x (GenBase.get_low W0 (S n) y)))
+ (fun n x y => reduce_10 (w10_modn1 (S n) x y))
+ (fun x y => reduce_11 (w11_mod_gt x y))
+ (fun n x y => reduce_11 (w11_mod_gt x (GenBase.get_low W0 (S n) y)))
+ (fun n x y => reduce_11 (w11_modn1 (S n) x y))
+ (fun x y => reduce_12 (w12_mod_gt x y))
+ (fun n x y => reduce_12 (w12_mod_gt x (GenBase.get_low W0 (S n) y)))
+ (fun n x y => reduce_12 (w12_modn1 (S n) x y))
+ (fun x y => reduce_13 (w13_mod_gt x y))
+ (fun n x y => reduce_13 (w13_mod_gt x (GenBase.get_low W0 (S n) y)))
+ (fun n x y => reduce_13 (w13_modn1 (S n) x y))
+ mod_gtnm).
+
+ Theorem spec_mod_gt:
+ forall x y, [x] > [y] -> 0 < [y] -> [mod_gt x y] = [x] mod [y].
+ Admitted.
Definition modulo x y :=
match compare x y with
@@ -3798,6 +3064,16 @@ Module Make (W0:W0Type).
| Gt => mod_gt x y
end.
+ Theorem spec_modulo:
+ forall x y, 0 < [y] -> [modulo x y] = [x] mod [y].
+ Admitted.
+
+ (***************************************************************)
+ (* *)
+ (* Gcd *)
+ (* *)
+ (***************************************************************)
+
Definition digits x :=
match x with
| N0 _ => w0_op.(znz_digits)
@@ -3817,6 +3093,9 @@ Module Make (W0:W0Type).
| Nn n _ => (make_op n).(znz_digits)
end.
+ Theorem spec_digits: forall x, 0 <= [x] < 2 ^ Zpos (digits x).
+ Admitted.
+
Definition gcd_gt_body a b cont :=
match compare b zero with
| Gt =>
@@ -3828,13 +3107,13 @@ Module Make (W0:W0Type).
| _ => a
end.
- Fixpoint gcd_gt (p:positive) (cont:t->t->t) (a b:t) {struct p} : t :=
+ Fixpoint gcd_gt_aux (p:positive) (cont:t->t->t) (a b:t) {struct p} : t :=
gcd_gt_body a b
(fun a b =>
match p with
| xH => cont a b
- | xO p => gcd_gt p (gcd_gt p cont) a b
- | xI p => gcd_gt p (gcd_gt p cont) a b
+ | xO p => gcd_gt_aux p (gcd_gt_aux p cont) a b
+ | xI p => gcd_gt_aux p (gcd_gt_aux p cont) a b
end).
Definition gcd_cont a b :=
@@ -3843,60 +3122,93 @@ Module Make (W0:W0Type).
| _ => a
end.
+ Definition gcd_gt a b := gcd_gt_aux (digits a) gcd_cont a b.
+
+ Theorem spec_gcd_gt: forall a b,
+ [a] > [b] -> [gcd_gt a b] = Zgcd [a] [b].
+ Admitted.
+
Definition gcd a b :=
match compare a b with
| Eq => a
- | Lt => gcd_gt (digits b) gcd_cont b a
- | Gt => gcd_gt (digits a) gcd_cont a b
+ | Lt => gcd_gt b a
+ | Gt => gcd_gt a b
end.
-Definition pheight p := Peano.pred (nat_of_P (get_height w0_op.(znz_digits) (plength p))).
+ Theorem spec_gcd: forall a b, [gcd a b] = Zgcd [a] [b].
+ Admitted.
+
+ (***************************************************************)
+ (* *)
+ (* Conversion *)
+ (* *)
+ (***************************************************************)
+
+ Definition pheight p :=
+ Peano.pred (nat_of_P (get_height w0_op.(znz_digits) (plength p))).
+
+ Theorem pheight_correct: forall p,
+ Zpos p < 2 ^ (Zpos (znz_digits w0_op) * 2 ^ (Z_of_nat (pheight p))).
+ Proof.
+ intros p; unfold pheight.
+ assert (F1: forall x, Z_of_nat (Peano.pred (nat_of_P x)) = Zpos x - 1).
+ intros x.
+ assert (Zsucc (Z_of_nat (Peano.pred (nat_of_P x))) = Zpos x); auto with zarith.
+ rewrite <- inj_S.
+ rewrite <- (fun x => S_pred x 0); auto with zarith.
+ rewrite Zpos_eq_Z_of_nat_o_nat_of_P; auto.
+ apply lt_le_trans with 1%nat; auto with zarith.
+ exact (le_Pmult_nat x 1).
+ rewrite F1; clear F1.
+ assert (F2:= (get_height_correct (znz_digits w0_op) (plength p))).
+ apply Zlt_le_trans with (Zpos (Psucc p)).
+ rewrite Zpos_succ_morphism; auto with zarith.
+ apply Zle_trans with (1 := plength_pred_correct (Psucc p)).
+ rewrite Ppred_succ.
+ apply Zpower_le_monotone; auto with zarith.
+ Qed.
+
Definition of_pos x :=
let h := pheight x in
match h with
- | O => reduce_0 (snd (w0_op.(znz_of_pos) x))
- | (S O) => reduce_1 (snd (w1_op.(znz_of_pos) x))
- | (S (S O)) => reduce_2 (snd (w2_op.(znz_of_pos) x))
- | (S (S (S O))) => reduce_3 (snd (w3_op.(znz_of_pos) x))
- | (S (S (S (S O)))) => reduce_4 (snd (w4_op.(znz_of_pos) x))
- | (S (S (S (S (S O))))) => reduce_5 (snd (w5_op.(znz_of_pos) x))
- | (S (S (S (S (S (S O)))))) => reduce_6 (snd (w6_op.(znz_of_pos) x))
- | (S (S (S (S (S (S (S O))))))) => reduce_7 (snd (w7_op.(znz_of_pos) x))
- | (S (S (S (S (S (S (S (S O)))))))) => reduce_8 (snd (w8_op.(znz_of_pos) x))
- | (S (S (S (S (S (S (S (S (S O))))))))) => reduce_9 (snd (w9_op.(znz_of_pos) x))
- | (S (S (S (S (S (S (S (S (S (S O)))))))))) => reduce_10 (snd (w10_op.(znz_of_pos) x))
- | (S (S (S (S (S (S (S (S (S (S (S O))))))))))) => reduce_11 (snd (w11_op.(znz_of_pos) x))
- | (S (S (S (S (S (S (S (S (S (S (S (S O)))))))))))) => reduce_12 (snd (w12_op.(znz_of_pos) x))
- | (S (S (S (S (S (S (S (S (S (S (S (S (S O))))))))))))) => reduce_13 (snd (w13_op.(znz_of_pos) x))
+ | 0%nat => reduce_0 (snd (w0_op.(znz_of_pos) x))
+ | 1%nat => reduce_1 (snd (w1_op.(znz_of_pos) x))
+ | 2%nat => reduce_2 (snd (w2_op.(znz_of_pos) x))
+ | 3%nat => reduce_3 (snd (w3_op.(znz_of_pos) x))
+ | 4%nat => reduce_4 (snd (w4_op.(znz_of_pos) x))
+ | 5%nat => reduce_5 (snd (w5_op.(znz_of_pos) x))
+ | 6%nat => reduce_6 (snd (w6_op.(znz_of_pos) x))
+ | 7%nat => reduce_7 (snd (w7_op.(znz_of_pos) x))
+ | 8%nat => reduce_8 (snd (w8_op.(znz_of_pos) x))
+ | 9%nat => reduce_9 (snd (w9_op.(znz_of_pos) x))
+ | 10%nat => reduce_10 (snd (w10_op.(znz_of_pos) x))
+ | 11%nat => reduce_11 (snd (w11_op.(znz_of_pos) x))
+ | 12%nat => reduce_12 (snd (w12_op.(znz_of_pos) x))
+ | 13%nat => reduce_13 (snd (w13_op.(znz_of_pos) x))
| _ =>
let n := minus h 14 in
reduce_n n (snd ((make_op n).(znz_of_pos) x))
end.
+ Theorem spec_of_pos: forall x,
+ [of_pos x] = Zpos x.
+ Admitted.
+
Definition of_N x :=
match x with
| BinNat.N0 => zero
| Npos p => of_pos p
end.
- Definition to_Z x :=
- match x with
- | N0 wx => w0_op.(znz_to_Z) wx
- | N1 wx => w1_op.(znz_to_Z) wx
- | N2 wx => w2_op.(znz_to_Z) wx
- | N3 wx => w3_op.(znz_to_Z) wx
- | N4 wx => w4_op.(znz_to_Z) wx
- | N5 wx => w5_op.(znz_to_Z) wx
- | N6 wx => w6_op.(znz_to_Z) wx
- | N7 wx => w7_op.(znz_to_Z) wx
- | N8 wx => w8_op.(znz_to_Z) wx
- | N9 wx => w9_op.(znz_to_Z) wx
- | N10 wx => w10_op.(znz_to_Z) wx
- | N11 wx => w11_op.(znz_to_Z) wx
- | N12 wx => w12_op.(znz_to_Z) wx
- | N13 wx => w13_op.(znz_to_Z) wx
- | Nn n wx => (make_op n).(znz_to_Z) wx
- end.
+ Theorem spec_of_N: forall x,
+ [of_N x] = Z_of_N x.
+ Admitted.
+
+ (***************************************************************)
+ (* *)
+ (* Shift *)
+ (* *)
+ (***************************************************************)
Definition head0 w := match w with
| N0 w=> reduce_0 (w0_op.(znz_head0) w)
@@ -3916,6 +3228,13 @@ Definition pheight p := Peano.pred (nat_of_P (get_height w0_op.(znz_digits) (ple
| Nn n w=> reduce_n n ((make_op n).(znz_head0) w)
end.
+ Theorem spec_head00: forall x, [x] = 0 ->[head0 x] = Zpos (digits x).
+ Admitted.
+
+ Theorem spec_head0: forall x, 0 < [x] ->
+ 2 ^ (Zpos (digits x) - 1) <= 2 ^ [head0 x] * [x] < 2 ^ Zpos (digits x).
+ Admitted.
+
Definition tail0 w := match w with
| N0 w=> reduce_0 (w0_op.(znz_tail0) w)
| N1 w=> reduce_1 (w1_op.(znz_tail0) w)
@@ -3934,6 +3253,13 @@ Definition pheight p := Peano.pred (nat_of_P (get_height w0_op.(znz_digits) (ple
| Nn n w=> reduce_n n ((make_op n).(znz_tail0) w)
end.
+ Theorem spec_tail00: forall x, [x] = 0 ->[tail0 x] = Zpos (digits x).
+ Admitted.
+
+ Theorem spec_tail0: forall x,
+ 0 < [x] -> exists y, 0 <= y /\ [x] = (2 * y + 1) * 2 ^ [tail0 x].
+ Admitted.
+
Definition Ndigits x :=
match x with
| N0 _ => N0 w0_op.(znz_zdigits)
@@ -3951,256 +3277,10 @@ Definition pheight p := Peano.pred (nat_of_P (get_height w0_op.(znz_digits) (ple
| N12 _ => reduce_12 w12_op.(znz_zdigits)
| N13 _ => reduce_13 w13_op.(znz_zdigits)
| Nn n _ => reduce_n n (make_op n).(znz_zdigits)
-
end.
- Definition level f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 fn x y: t_ := match x, y with
- | N0 wx, N0 wy => f0 wx wy
- | N0 wx, N1 wy => f1 (WW w_0 wx) wy
- | N0 wx, N2 wy => f2 (extend1 w0 (WW w_0 wx)) wy
- | N0 wx, N3 wy => f3 (extend2 w0 (WW w_0 wx)) wy
- | N0 wx, N4 wy => f4 (extend3 w0 (WW w_0 wx)) wy
- | N0 wx, N5 wy => f5 (extend4 w0 (WW w_0 wx)) wy
- | N0 wx, N6 wy => f6 (extend5 w0 (WW w_0 wx)) wy
- | N0 wx, N7 wy => f7 (extend6 w0 (WW w_0 wx)) wy
- | N0 wx, N8 wy => f8 (extend7 w0 (WW w_0 wx)) wy
- | N0 wx, N9 wy => f9 (extend8 w0 (WW w_0 wx)) wy
- | N0 wx, N10 wy => f10 (extend9 w0 (WW w_0 wx)) wy
- | N0 wx, N11 wy => f11 (extend10 w0 (WW w_0 wx)) wy
- | N0 wx, N12 wy => f12 (extend11 w0 (WW w_0 wx)) wy
- | N0 wx, N13 wy => f13 (extend12 w0 (WW w_0 wx)) wy
- | N0 wx, Nn n wy =>
- fn n (extend n w13 (extend13 w0 (WW w_0 wx))) wy
- | N1 wx, N0 wy =>
- f1 wx (WW w_0 wy)
- | N1 wx, N1 wy => f1 wx wy
- | N1 wx, N2 wy => f2 (extend1 w0 wx) wy
- | N1 wx, N3 wy => f3 (extend2 w0 wx) wy
- | N1 wx, N4 wy => f4 (extend3 w0 wx) wy
- | N1 wx, N5 wy => f5 (extend4 w0 wx) wy
- | N1 wx, N6 wy => f6 (extend5 w0 wx) wy
- | N1 wx, N7 wy => f7 (extend6 w0 wx) wy
- | N1 wx, N8 wy => f8 (extend7 w0 wx) wy
- | N1 wx, N9 wy => f9 (extend8 w0 wx) wy
- | N1 wx, N10 wy => f10 (extend9 w0 wx) wy
- | N1 wx, N11 wy => f11 (extend10 w0 wx) wy
- | N1 wx, N12 wy => f12 (extend11 w0 wx) wy
- | N1 wx, N13 wy => f13 (extend12 w0 wx) wy
- | N1 wx, Nn n wy => fn n (extend n w13 (extend13 w0 wx)) wy
- | N2 wx, N0 wy =>
- f2 wx (extend1 w0 (WW w_0 wy))
- | N2 wx, N1 wy => f2 wx (extend1 w0 wy)
- | N2 wx, N2 wy => f2 wx wy
- | N2 wx, N3 wy => f3 (extend1 w1 wx) wy
- | N2 wx, N4 wy => f4 (extend2 w1 wx) wy
- | N2 wx, N5 wy => f5 (extend3 w1 wx) wy
- | N2 wx, N6 wy => f6 (extend4 w1 wx) wy
- | N2 wx, N7 wy => f7 (extend5 w1 wx) wy
- | N2 wx, N8 wy => f8 (extend6 w1 wx) wy
- | N2 wx, N9 wy => f9 (extend7 w1 wx) wy
- | N2 wx, N10 wy => f10 (extend8 w1 wx) wy
- | N2 wx, N11 wy => f11 (extend9 w1 wx) wy
- | N2 wx, N12 wy => f12 (extend10 w1 wx) wy
- | N2 wx, N13 wy => f13 (extend11 w1 wx) wy
- | N2 wx, Nn n wy => fn n (extend n w13 (extend12 w1 wx)) wy
- | N3 wx, N0 wy =>
- f3 wx (extend2 w0 (WW w_0 wy))
- | N3 wx, N1 wy => f3 wx (extend2 w0 wy)
- | N3 wx, N2 wy => f3 wx (extend1 w1 wy)
- | N3 wx, N3 wy => f3 wx wy
- | N3 wx, N4 wy => f4 (extend1 w2 wx) wy
- | N3 wx, N5 wy => f5 (extend2 w2 wx) wy
- | N3 wx, N6 wy => f6 (extend3 w2 wx) wy
- | N3 wx, N7 wy => f7 (extend4 w2 wx) wy
- | N3 wx, N8 wy => f8 (extend5 w2 wx) wy
- | N3 wx, N9 wy => f9 (extend6 w2 wx) wy
- | N3 wx, N10 wy => f10 (extend7 w2 wx) wy
- | N3 wx, N11 wy => f11 (extend8 w2 wx) wy
- | N3 wx, N12 wy => f12 (extend9 w2 wx) wy
- | N3 wx, N13 wy => f13 (extend10 w2 wx) wy
- | N3 wx, Nn n wy => fn n (extend n w13 (extend11 w2 wx)) wy
- | N4 wx, N0 wy =>
- f4 wx (extend3 w0 (WW w_0 wy))
- | N4 wx, N1 wy => f4 wx (extend3 w0 wy)
- | N4 wx, N2 wy => f4 wx (extend2 w1 wy)
- | N4 wx, N3 wy => f4 wx (extend1 w2 wy)
- | N4 wx, N4 wy => f4 wx wy
- | N4 wx, N5 wy => f5 (extend1 w3 wx) wy
- | N4 wx, N6 wy => f6 (extend2 w3 wx) wy
- | N4 wx, N7 wy => f7 (extend3 w3 wx) wy
- | N4 wx, N8 wy => f8 (extend4 w3 wx) wy
- | N4 wx, N9 wy => f9 (extend5 w3 wx) wy
- | N4 wx, N10 wy => f10 (extend6 w3 wx) wy
- | N4 wx, N11 wy => f11 (extend7 w3 wx) wy
- | N4 wx, N12 wy => f12 (extend8 w3 wx) wy
- | N4 wx, N13 wy => f13 (extend9 w3 wx) wy
- | N4 wx, Nn n wy => fn n (extend n w13 (extend10 w3 wx)) wy
- | N5 wx, N0 wy =>
- f5 wx (extend4 w0 (WW w_0 wy))
- | N5 wx, N1 wy => f5 wx (extend4 w0 wy)
- | N5 wx, N2 wy => f5 wx (extend3 w1 wy)
- | N5 wx, N3 wy => f5 wx (extend2 w2 wy)
- | N5 wx, N4 wy => f5 wx (extend1 w3 wy)
- | N5 wx, N5 wy => f5 wx wy
- | N5 wx, N6 wy => f6 (extend1 w4 wx) wy
- | N5 wx, N7 wy => f7 (extend2 w4 wx) wy
- | N5 wx, N8 wy => f8 (extend3 w4 wx) wy
- | N5 wx, N9 wy => f9 (extend4 w4 wx) wy
- | N5 wx, N10 wy => f10 (extend5 w4 wx) wy
- | N5 wx, N11 wy => f11 (extend6 w4 wx) wy
- | N5 wx, N12 wy => f12 (extend7 w4 wx) wy
- | N5 wx, N13 wy => f13 (extend8 w4 wx) wy
- | N5 wx, Nn n wy => fn n (extend n w13 (extend9 w4 wx)) wy
- | N6 wx, N0 wy =>
- f6 wx (extend5 w0 (WW w_0 wy))
- | N6 wx, N1 wy => f6 wx (extend5 w0 wy)
- | N6 wx, N2 wy => f6 wx (extend4 w1 wy)
- | N6 wx, N3 wy => f6 wx (extend3 w2 wy)
- | N6 wx, N4 wy => f6 wx (extend2 w3 wy)
- | N6 wx, N5 wy => f6 wx (extend1 w4 wy)
- | N6 wx, N6 wy => f6 wx wy
- | N6 wx, N7 wy => f7 (extend1 w5 wx) wy
- | N6 wx, N8 wy => f8 (extend2 w5 wx) wy
- | N6 wx, N9 wy => f9 (extend3 w5 wx) wy
- | N6 wx, N10 wy => f10 (extend4 w5 wx) wy
- | N6 wx, N11 wy => f11 (extend5 w5 wx) wy
- | N6 wx, N12 wy => f12 (extend6 w5 wx) wy
- | N6 wx, N13 wy => f13 (extend7 w5 wx) wy
- | N6 wx, Nn n wy => fn n (extend n w13 (extend8 w5 wx)) wy
- | N7 wx, N0 wy =>
- f7 wx (extend6 w0 (WW w_0 wy))
- | N7 wx, N1 wy => f7 wx (extend6 w0 wy)
- | N7 wx, N2 wy => f7 wx (extend5 w1 wy)
- | N7 wx, N3 wy => f7 wx (extend4 w2 wy)
- | N7 wx, N4 wy => f7 wx (extend3 w3 wy)
- | N7 wx, N5 wy => f7 wx (extend2 w4 wy)
- | N7 wx, N6 wy => f7 wx (extend1 w5 wy)
- | N7 wx, N7 wy => f7 wx wy
- | N7 wx, N8 wy => f8 (extend1 w6 wx) wy
- | N7 wx, N9 wy => f9 (extend2 w6 wx) wy
- | N7 wx, N10 wy => f10 (extend3 w6 wx) wy
- | N7 wx, N11 wy => f11 (extend4 w6 wx) wy
- | N7 wx, N12 wy => f12 (extend5 w6 wx) wy
- | N7 wx, N13 wy => f13 (extend6 w6 wx) wy
- | N7 wx, Nn n wy => fn n (extend n w13 (extend7 w6 wx)) wy
- | N8 wx, N0 wy =>
- f8 wx (extend7 w0 (WW w_0 wy))
- | N8 wx, N1 wy => f8 wx (extend7 w0 wy)
- | N8 wx, N2 wy => f8 wx (extend6 w1 wy)
- | N8 wx, N3 wy => f8 wx (extend5 w2 wy)
- | N8 wx, N4 wy => f8 wx (extend4 w3 wy)
- | N8 wx, N5 wy => f8 wx (extend3 w4 wy)
- | N8 wx, N6 wy => f8 wx (extend2 w5 wy)
- | N8 wx, N7 wy => f8 wx (extend1 w6 wy)
- | N8 wx, N8 wy => f8 wx wy
- | N8 wx, N9 wy => f9 (extend1 w7 wx) wy
- | N8 wx, N10 wy => f10 (extend2 w7 wx) wy
- | N8 wx, N11 wy => f11 (extend3 w7 wx) wy
- | N8 wx, N12 wy => f12 (extend4 w7 wx) wy
- | N8 wx, N13 wy => f13 (extend5 w7 wx) wy
- | N8 wx, Nn n wy => fn n (extend n w13 (extend6 w7 wx)) wy
- | N9 wx, N0 wy =>
- f9 wx (extend8 w0 (WW w_0 wy))
- | N9 wx, N1 wy => f9 wx (extend8 w0 wy)
- | N9 wx, N2 wy => f9 wx (extend7 w1 wy)
- | N9 wx, N3 wy => f9 wx (extend6 w2 wy)
- | N9 wx, N4 wy => f9 wx (extend5 w3 wy)
- | N9 wx, N5 wy => f9 wx (extend4 w4 wy)
- | N9 wx, N6 wy => f9 wx (extend3 w5 wy)
- | N9 wx, N7 wy => f9 wx (extend2 w6 wy)
- | N9 wx, N8 wy => f9 wx (extend1 w7 wy)
- | N9 wx, N9 wy => f9 wx wy
- | N9 wx, N10 wy => f10 (extend1 w8 wx) wy
- | N9 wx, N11 wy => f11 (extend2 w8 wx) wy
- | N9 wx, N12 wy => f12 (extend3 w8 wx) wy
- | N9 wx, N13 wy => f13 (extend4 w8 wx) wy
- | N9 wx, Nn n wy => fn n (extend n w13 (extend5 w8 wx)) wy
- | N10 wx, N0 wy =>
- f10 wx (extend9 w0 (WW w_0 wy))
- | N10 wx, N1 wy => f10 wx (extend9 w0 wy)
- | N10 wx, N2 wy => f10 wx (extend8 w1 wy)
- | N10 wx, N3 wy => f10 wx (extend7 w2 wy)
- | N10 wx, N4 wy => f10 wx (extend6 w3 wy)
- | N10 wx, N5 wy => f10 wx (extend5 w4 wy)
- | N10 wx, N6 wy => f10 wx (extend4 w5 wy)
- | N10 wx, N7 wy => f10 wx (extend3 w6 wy)
- | N10 wx, N8 wy => f10 wx (extend2 w7 wy)
- | N10 wx, N9 wy => f10 wx (extend1 w8 wy)
- | N10 wx, N10 wy => f10 wx wy
- | N10 wx, N11 wy => f11 (extend1 w9 wx) wy
- | N10 wx, N12 wy => f12 (extend2 w9 wx) wy
- | N10 wx, N13 wy => f13 (extend3 w9 wx) wy
- | N10 wx, Nn n wy => fn n (extend n w13 (extend4 w9 wx)) wy
- | N11 wx, N0 wy =>
- f11 wx (extend10 w0 (WW w_0 wy))
- | N11 wx, N1 wy => f11 wx (extend10 w0 wy)
- | N11 wx, N2 wy => f11 wx (extend9 w1 wy)
- | N11 wx, N3 wy => f11 wx (extend8 w2 wy)
- | N11 wx, N4 wy => f11 wx (extend7 w3 wy)
- | N11 wx, N5 wy => f11 wx (extend6 w4 wy)
- | N11 wx, N6 wy => f11 wx (extend5 w5 wy)
- | N11 wx, N7 wy => f11 wx (extend4 w6 wy)
- | N11 wx, N8 wy => f11 wx (extend3 w7 wy)
- | N11 wx, N9 wy => f11 wx (extend2 w8 wy)
- | N11 wx, N10 wy => f11 wx (extend1 w9 wy)
- | N11 wx, N11 wy => f11 wx wy
- | N11 wx, N12 wy => f12 (extend1 w10 wx) wy
- | N11 wx, N13 wy => f13 (extend2 w10 wx) wy
- | N11 wx, Nn n wy => fn n (extend n w13 (extend3 w10 wx)) wy
- | N12 wx, N0 wy =>
- f12 wx (extend11 w0 (WW w_0 wy))
- | N12 wx, N1 wy => f12 wx (extend11 w0 wy)
- | N12 wx, N2 wy => f12 wx (extend10 w1 wy)
- | N12 wx, N3 wy => f12 wx (extend9 w2 wy)
- | N12 wx, N4 wy => f12 wx (extend8 w3 wy)
- | N12 wx, N5 wy => f12 wx (extend7 w4 wy)
- | N12 wx, N6 wy => f12 wx (extend6 w5 wy)
- | N12 wx, N7 wy => f12 wx (extend5 w6 wy)
- | N12 wx, N8 wy => f12 wx (extend4 w7 wy)
- | N12 wx, N9 wy => f12 wx (extend3 w8 wy)
- | N12 wx, N10 wy => f12 wx (extend2 w9 wy)
- | N12 wx, N11 wy => f12 wx (extend1 w10 wy)
- | N12 wx, N12 wy => f12 wx wy
- | N12 wx, N13 wy => f13 (extend1 w11 wx) wy
- | N12 wx, Nn n wy => fn n (extend n w13 (extend2 w11 wx)) wy
- | N13 wx, N0 wy =>
- f13 wx (extend12 w0 (WW w_0 wy))
- | N13 wx, N1 wy => f13 wx (extend12 w0 wy)
- | N13 wx, N2 wy => f13 wx (extend11 w1 wy)
- | N13 wx, N3 wy => f13 wx (extend10 w2 wy)
- | N13 wx, N4 wy => f13 wx (extend9 w3 wy)
- | N13 wx, N5 wy => f13 wx (extend8 w4 wy)
- | N13 wx, N6 wy => f13 wx (extend7 w5 wy)
- | N13 wx, N7 wy => f13 wx (extend6 w6 wy)
- | N13 wx, N8 wy => f13 wx (extend5 w7 wy)
- | N13 wx, N9 wy => f13 wx (extend4 w8 wy)
- | N13 wx, N10 wy => f13 wx (extend3 w9 wy)
- | N13 wx, N11 wy => f13 wx (extend2 w10 wy)
- | N13 wx, N12 wy => f13 wx (extend1 w11 wy)
- | N13 wx, N13 wy => f13 wx wy
- | N13 wx, Nn n wy => fn n (extend n w13 (extend1 w12 wx)) wy
- | Nn n wx, N0 wy =>
- fn n wx (extend n w13 (extend13 w0 (WW w_0 wy)))
- | Nn n wx, N1 wy => fn n wx (extend n w13 (extend13 w0 wy))
- | Nn n wx, N2 wy => fn n wx (extend n w13 (extend12 w1 wy))
- | Nn n wx, N3 wy => fn n wx (extend n w13 (extend11 w2 wy))
- | Nn n wx, N4 wy => fn n wx (extend n w13 (extend10 w3 wy))
- | Nn n wx, N5 wy => fn n wx (extend n w13 (extend9 w4 wy))
- | Nn n wx, N6 wy => fn n wx (extend n w13 (extend8 w5 wy))
- | Nn n wx, N7 wy => fn n wx (extend n w13 (extend7 w6 wy))
- | Nn n wx, N8 wy => fn n wx (extend n w13 (extend6 w7 wy))
- | Nn n wx, N9 wy => fn n wx (extend n w13 (extend5 w8 wy))
- | Nn n wx, N10 wy => fn n wx (extend n w13 (extend4 w9 wy))
- | Nn n wx, N11 wy => fn n wx (extend n w13 (extend3 w10 wy))
- | Nn n wx, N12 wy => fn n wx (extend n w13 (extend2 w11 wy))
- | Nn n wx, N13 wy => fn n wx (extend n w13 (extend1 w12 wy))
- | Nn n wx, Nn m wy =>
- let mn := Max.max n m in
- let d := diff n m in
- fn mn
- (castm (diff_r n m) (extend_tr wx (snd d)))
- (castm (diff_l n m) (extend_tr wy (fst d)))
- end.
+ Theorem spec_Ndigits: forall x, [Ndigits x] = Zpos (digits x).
+ Admitted.
Definition shiftr0 n x := w0_op.(znz_add_mul_div) (w0_op.(znz_sub) w0_op.(znz_zdigits) n) w0_op.(znz_0) x.
Definition shiftr1 n x := w1_op.(znz_add_mul_div) (w1_op.(znz_sub) w1_op.(znz_zdigits) n) w1_op.(znz_0) x.
@@ -4218,9 +3298,8 @@ Definition pheight p := Peano.pred (nat_of_P (get_height w0_op.(znz_digits) (ple
Definition shiftr13 n x := w13_op.(znz_add_mul_div) (w13_op.(znz_sub) w13_op.(znz_zdigits) n) w13_op.(znz_0) x.
Definition shiftrn n p x := (make_op n).(znz_add_mul_div) ((make_op n).(znz_sub) (make_op n).(znz_zdigits) p) (make_op n).(znz_0) x.
- Definition shiftr :=
- Eval lazy beta delta [level] in
- level (fun n x => N0 (shiftr0 n x))
+ Definition shiftr := Eval lazy beta delta [same_level] in
+ same_level _ (fun n x => N0 (shiftr0 n x))
(fun n x => reduce_1 (shiftr1 n x))
(fun n x => reduce_2 (shiftr2 n x))
(fun n x => reduce_3 (shiftr3 n x))
@@ -4236,12 +3315,21 @@ Definition pheight p := Peano.pred (nat_of_P (get_height w0_op.(znz_digits) (ple
(fun n x => reduce_13 (shiftr13 n x))
(fun n p x => reduce_n n (shiftrn n p x)).
+ Theorem spec_shiftr: forall n x,
+ [n] <= [Ndigits x] -> [shiftr n x] = [x] / 2 ^ [n].
+ Admitted.
+
Definition safe_shiftr n x :=
match compare n (Ndigits x) with
| Lt => shiftr n x
| _ => N0 w_0
end.
+ Theorem spec_safe_shiftr: forall n x,
+ [safe_shiftr n x] = [x] / 2 ^ [n].
+ Admitted.
+
+
Definition shiftl0 n x := w0_op.(znz_add_mul_div) n x w0_op.(znz_0).
Definition shiftl1 n x := w1_op.(znz_add_mul_div) n x w1_op.(znz_0).
Definition shiftl2 n x := w2_op.(znz_add_mul_div) n x w2_op.(znz_0).
@@ -4257,9 +3345,8 @@ Definition pheight p := Peano.pred (nat_of_P (get_height w0_op.(znz_digits) (ple
Definition shiftl12 n x := w12_op.(znz_add_mul_div) n x w12_op.(znz_0).
Definition shiftl13 n x := w13_op.(znz_add_mul_div) n x w13_op.(znz_0).
Definition shiftln n p x := (make_op n).(znz_add_mul_div) p x (make_op n).(znz_0).
- Definition shiftl :=
- Eval lazy beta delta [level] in
- level (fun n x => N0 (shiftl0 n x))
+ Definition shiftl := Eval lazy beta delta [same_level] in
+ same_level _ (fun n x => N0 (shiftl0 n x))
(fun n x => reduce_1 (shiftl1 n x))
(fun n x => reduce_2 (shiftl2 n x))
(fun n x => reduce_3 (shiftl3 n x))
@@ -4276,30 +3363,56 @@ Definition pheight p := Peano.pred (nat_of_P (get_height w0_op.(znz_digits) (ple
(fun n p x => reduce_n n (shiftln n p x)).
+ Theorem spec_shiftl: forall n x,
+ [n] <= [head0 x] -> [shiftl n x] = [x] * 2 ^ [n].
+ Admitted.
+
Definition double_size w := match w with
- | N0 w=> N1 (WW w_0 w)
- | N1 w=> N2 (extend1 _ w)
- | N2 w=> N3 (extend1 _ w)
- | N3 w=> N4 (extend1 _ w)
- | N4 w=> N5 (extend1 _ w)
- | N5 w=> N6 (extend1 _ w)
- | N6 w=> N7 (extend1 _ w)
- | N7 w=> N8 (extend1 _ w)
- | N8 w=> N9 (extend1 _ w)
- | N9 w=> N10 (extend1 _ w)
- | N10 w=> N11 (extend1 _ w)
- | N11 w=> N12 (extend1 _ w)
- | N12 w=> N13 (extend1 _ w)
- | N13 w=> Nn 0 (extend1 _ w)
- | Nn n w=> Nn (S n) (extend1 _ w)
+ | N0 x => N1 (WW (znz_0 w0_op) x)
+ | N1 x => N2 (WW (znz_0 w1_op) x)
+ | N2 x => N3 (WW (znz_0 w2_op) x)
+ | N3 x => N4 (WW (znz_0 w3_op) x)
+ | N4 x => N5 (WW (znz_0 w4_op) x)
+ | N5 x => N6 (WW (znz_0 w5_op) x)
+ | N6 x => N7 (WW (znz_0 w6_op) x)
+ | N7 x => N8 (WW (znz_0 w7_op) x)
+ | N8 x => N9 (WW (znz_0 w8_op) x)
+ | N9 x => N10 (WW (znz_0 w9_op) x)
+ | N10 x => N11 (WW (znz_0 w10_op) x)
+ | N11 x => N12 (WW (znz_0 w11_op) x)
+ | N12 x => N13 (WW (znz_0 w12_op) x)
+ | N13 x => Nn 0 (WW (znz_0 w13_op) x)
+ | Nn n x => Nn (S n) (WW (znz_0 (make_op n)) x)
end.
+ Theorem spec_double_size_digits:
+ forall x, digits (double_size x) = xO (digits x).
+ Admitted.
+
+ Theorem spec_double_size: forall x, [double_size x] = [x].
+ Admitted.
+
+ Theorem spec_double_size_head0:
+ forall x, 2 * [head0 x] <= [head0 (double_size x)].
+ Admitted.
+
+ Theorem spec_double_size_head0_pos:
+ forall x, 0 < [head0 (double_size x)].
+ Admitted.
+
Definition safe_shiftl_aux_body cont n x :=
match compare n (head0 x) with
Gt => cont n (double_size x)
| _ => shiftl n x
end.
+ Theorem spec_safe_shift_aux_body: forall n p x cont,
+ 2^ Zpos p <= [head0 x] ->
+ (forall x, 2 ^ (Zpos p + 1) <= [head0 x]->
+ [cont n x] = [x] * 2 ^ [n]) ->
+ [safe_shiftl_aux_body cont n x] = [x] * 2 ^ [n].
+ Admitted.
+
Fixpoint safe_shiftl_aux p cont n x {struct p} :=
safe_shiftl_aux_body
(fun n x => match p with
@@ -4308,9 +3421,22 @@ Definition pheight p := Peano.pred (nat_of_P (get_height w0_op.(znz_digits) (ple
| xI p => safe_shiftl_aux p (safe_shiftl_aux p cont) n x
end) n x.
- Definition safe_shiftl n x :=
- safe_shiftl_aux (digits n) (fun n x => N0 w0_op.(znz_0)) n x.
-
+ Theorem spec_safe_shift_aux: forall p q n x cont,
+ 2 ^ (Zpos q) <= [head0 x] ->
+ (forall x, 2 ^ (Zpos p + Zpos q) <= [head0 x] ->
+ [cont n x] = [x] * 2 ^ [n]) ->
+ [safe_shiftl_aux p cont n x] = [x] * 2 ^ [n].
+ Admitted.
+
+ Definition safe_shiftl n x :=
+ safe_shiftl_aux_body
+ (safe_shiftl_aux_body
+ (safe_shiftl_aux (digits n) shiftl)) n x.
+
+ Theorem spec_safe_shift: forall n x,
+ [safe_shiftl n x] = [x] * 2 ^ [n].
+ Admitted.
+
Definition is_even x :=
match x with
| N0 wx => w0_op.(znz_is_even) wx
@@ -4330,39 +3456,8 @@ Definition pheight p := Peano.pred (nat_of_P (get_height w0_op.(znz_digits) (ple
| Nn n wx => (make_op n).(znz_is_even) wx
end.
-(* Proof section *)
-
- Open Scope Z_scope.
- Notation "[ x ]" := (to_Z x).
-
- Theorem succ_spec: forall n, [succ n] = [n] + 1.
- Admitted.
-
- Theorem spec_add: forall x y, [add x y] = [x] + [y].
- Admitted.
-
- Theorem spec_sub: forall x y, [y] <= [x] -> [sub x y] = [x] - [y].
- Admitted.
-
- Theorem spec_sub0: forall x y, [x] < [y] -> [sub x y] = 0.
- Admitted.
-
- Theorem spec_compare: forall x y,
- match compare x y with
- Eq => [x] = [y]
- | Lt => [x] < [y]
- | Gt => [x] > [y]
- end.
- Proof.
- Admitted.
-
- Theorem spec_mul: forall x y, [mul x y] =[x] * [y].
- Proof.
- Admitted.
-
- Theorem spec_sqrt : forall x,
- [sqrt x] ^ 2 <= [x] < ([sqrt x] + 1) ^ 2.
- Proof.
+ Theorem spec_is_even: forall x,
+ if is_even x then [x] mod 2 = 0 else [x] mod 2 = 1.
Admitted.
End Make.
diff --git a/theories/Ints/num/Zn2Z.v b/theories/Ints/num/Zn2Z.v
index 2925b202d..db3f28b41 100644
--- a/theories/Ints/num/Zn2Z.v
+++ b/theories/Ints/num/Zn2Z.v
@@ -550,12 +550,11 @@ Section Zn2Z.
exact (spec_W0 op_spec). exact (spec_mul_c op_spec).
Qed.
-Axiom ok:forall P, P.
Let spec_ww_karatsuba_c : forall x y, [[karatsuba_c x y ]] = [|x|] * [|y|].
Proof.
refine (spec_ww_karatsuba_c _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _
_ _ _ _ _ _ _ _ _ _ _ _); auto.
-apply ok.
+ unfold w_digits; apply spec_more_than_1_digit; auto.
exact (spec_WW op_spec).
exact (spec_W0 op_spec).
exact (spec_compare op_spec).
@@ -641,6 +640,17 @@ apply ok.
rewrite Zpos_xO; auto with zarith.
Qed.
+
+ Let spec_ww_head00 : forall x, [|x|] = 0 -> [|head0 x|] = Zpos _ww_digits.
+ Proof.
+ refine (spec_ww_head00 w_0 w_0W
+ w_compare w_head0 w_add2 w_zdigits _ww_zdigits
+ w_to_Z _ _ _ (refl_equal _ww_digits) _ _ _ _); auto.
+ exact (spec_compare op_spec).
+ exact (spec_head00 op_spec).
+ exact (spec_zdigits op_spec).
+ Qed.
+
Let spec_ww_head0 : forall x, 0 < [|x|] ->
wwB/ 2 <= 2 ^ [|head0 x|] * [|x|] < wwB.
Proof.
@@ -652,6 +662,17 @@ apply ok.
exact (spec_zdigits op_spec).
Qed.
+ Let spec_ww_tail00 : forall x, [|x|] = 0 -> [|tail0 x|] = Zpos _ww_digits.
+ Proof.
+ refine (spec_ww_tail00 w_0 w_0W
+ w_compare w_tail0 w_add2 w_zdigits _ww_zdigits
+ w_to_Z _ _ _ (refl_equal _ww_digits) _ _ _ _); auto.
+ exact (spec_compare op_spec).
+ exact (spec_tail00 op_spec).
+ exact (spec_zdigits op_spec).
+ Qed.
+
+
Let spec_ww_tail0 : forall x, 0 < [|x|] ->
exists y, 0 <= y /\ [|x|] = (2 * y + 1) * 2 ^ [|tail0 x|].
Proof.
diff --git a/theories/Ints/num/ZnZ.v b/theories/Ints/num/ZnZ.v
index 1e0317031..68f1092cd 100644
--- a/theories/Ints/num/ZnZ.v
+++ b/theories/Ints/num/ZnZ.v
@@ -239,8 +239,10 @@ Section Spec.
(* shift operations *)
+ spec_head00: forall x, [|x|] = 0 -> [|w_head0 x|] = Zpos w_digits;
spec_head0 : forall x, 0 < [|x|] ->
wB/ 2 <= 2 ^ ([|w_head0 x|]) * [|x|] < wB;
+ spec_tail00: forall x, [|x|] = 0 -> [|w_tail0 x|] = Zpos w_digits;
spec_tail0 : forall x, 0 < [|x|] ->
exists y, 0 <= y /\ [|x|] = (2 * y + 1) * (2 ^ [|w_tail0 x|]) ;
spec_add_mul_div : forall x y p,
diff --git a/theories/Ints/num/genN.ml b/theories/Ints/num/genN.ml
index a2632f493..928a60402 100644
--- a/theories/Ints/num/genN.ml
+++ b/theories/Ints/num/genN.ml
@@ -2,24 +2,15 @@ open Format
let size = 13
let sizeaux = 1
+let gen_proof = false
let t = "t"
let c = "N"
-let gen_proof = false
-let gen_proof1 = (* true *) false
-let gen_proof2 = (* true *) false
-let gen_proof3 = false
-let gen_proof4 = (* true *) false
-let gen_proof5 = false
-let gen_proof6 = (* true *) false
-let gen_proof7 = false
-let gen_proof8 = false
-let gen_proof9 = false
-let gen_proof10 = (* true *) false
-let gen_proof11 = false
-let gen_proof12 = false
-let gen_proof13 = false
-let gen_proof14 = (* true *) false
+let pz n = if n == 0 then "w_0" else "W0"
+let rec gen2 n = if n == 0 then "1" else if n == 1 then "2"
+ else "2 * " ^ (gen2 (n - 1))
+let rec genxO n s =
+ if n == 0 then s else " (xO" ^ (genxO (n - 1) s) ^ ")"
(******* Start Printing ********)
@@ -27,7 +18,7 @@ let basename = "N"
let print_header fmt l =
- let l = "ZArith"::"Basic_type"::"ZnZ"::"Zn2Z"::"Nbasic"::"GenMul"::
+ let l = "ZAux"::"ZArith"::"Basic_type"::"ZnZ"::"Zn2Z"::"Nbasic"::"GenMul"::
"GenDivn1"::"Wf_nat"::l in
List.iter (fun s -> fprintf fmt "Require Import %s.\n" s) l;
fprintf fmt "\n"
@@ -53,6 +44,14 @@ let start_file post l =
let print_Make () =
let fmt = start_file "Make" [] in
+ fprintf fmt " (***************************************************************)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (* File automatically generated DO NOT EDIT *)\n";
+ fprintf fmt " (* Constructors: %i Generated Proofs: %b %s %s *)\n" size gen_proof (if size < 10 then " " else "") (if gen_proof then " " else "");
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (***************************************************************)\n\n";
+
+
fprintf fmt "Module Type W0Type.\n";
fprintf fmt " Parameter w : Set.\n";
fprintf fmt " Parameter w_op : znz_op w.\n";
@@ -120,152 +119,808 @@ let print_Make () =
done;
fprintf fmt "\n";
+
fprintf fmt " Definition zero := %s0 w_0.\n" c;
fprintf fmt " Definition one := %s0 one0.\n" c;
fprintf fmt "\n";
- (* Successor function *)
+ fprintf fmt " Definition to_Z x :=\n";
+ fprintf fmt " match x with\n";
for i = 0 to size do
- fprintf fmt " Definition w%i_succ_c := w%i_op.(znz_succ_c).\n" i i
+ fprintf fmt " | %s%i wx => w%i_op.(znz_to_Z) wx\n" c i i
done;
+ fprintf fmt " | %sn n wx => (make_op n).(znz_to_Z) wx\n" c;
+ fprintf fmt " end.\n";
fprintf fmt "\n";
+ fprintf fmt " Open Scope Z_scope.\n";
+ fprintf fmt " Notation \"[ x ]\" := (to_Z x).\n";
+ fprintf fmt " \n";
+
+
+ if gen_proof then
+ begin
+ fprintf fmt " (* Regular make op (no karatsuba) *)\n";
+ fprintf fmt " Fixpoint nmake_op (ww:Set) (ww_op: znz_op ww) (n: nat) : \n";
+ fprintf fmt " znz_op (word ww n) :=\n";
+ fprintf fmt " match n return znz_op (word ww n) with \n";
+ fprintf fmt " O => ww_op\n";
+ fprintf fmt " | S n1 => mk_zn2z_op (nmake_op ww ww_op n1) \n";
+ fprintf fmt " end.\n";
+ fprintf fmt "\n";
+ fprintf fmt " (* Simplification by rewriting for nmake_op *)\n";
+ fprintf fmt " Theorem nmake_op_S: forall ww (w_op: znz_op ww) x, \n";
+ fprintf fmt " nmake_op _ w_op (S x) = mk_zn2z_op (nmake_op _ w_op x).\n";
+ fprintf fmt " auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+ end;
+
+
+ fprintf fmt " (* Eval and extend functions for each level *)\n";
for i = 0 to size do
- fprintf fmt " Definition w%i_succ := w%i_op.(znz_succ).\n" i i
+ if gen_proof then
+ fprintf fmt " Let nmake_op%i := nmake_op _ w%i_op.\n" i i;
+ if gen_proof then
+ fprintf fmt " Let eval%in n := znz_to_Z (nmake_op%i n).\n" i i;
+ if i == 0 then
+ fprintf fmt " Let extend%i := GenBase.extend (WW w_0).\n" i
+ else
+ fprintf fmt " Let extend%i := GenBase.extend (WW (W0: w%i)).\n" i i;
done;
fprintf fmt "\n";
- fprintf fmt " Definition succ x :=\n";
- fprintf fmt " match x with\n";
- for i = 0 to size-1 do
- fprintf fmt " | %s%i wx =>\n" c i;
- fprintf fmt " match w%i_succ_c wx with\n" i;
- fprintf fmt " | C0 r => %s%i r\n" c i;
- fprintf fmt " | C1 r => %s%i (WW one%i r)\n" c (i+1) i;
- fprintf fmt " end\n";
+
+ if gen_proof then
+ begin
+ fprintf fmt " Theorem digits_gend:forall n ww (w_op: znz_op ww), \n";
+ fprintf fmt " znz_digits (nmake_op _ w_op n) = \n";
+ fprintf fmt " GenBase.gen_digits (znz_digits w_op) n.\n";
+ fprintf fmt " Proof.";
+ fprintf fmt " intros n; elim n; auto; clear n.\n";
+ fprintf fmt " intros n Hrec ww ww_op; simpl GenBase.gen_digits.\n";
+ fprintf fmt " rewrite <- Hrec; auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+ fprintf fmt " Theorem nmake_gen: forall n ww (w_op: znz_op ww), \n";
+ fprintf fmt " znz_to_Z (nmake_op _ w_op n) =\n";
+ fprintf fmt " %sGenBase.gen_to_Z _ (znz_digits w_op) (znz_to_Z w_op) n.\n" "@";
+ fprintf fmt " Proof.";
+ fprintf fmt " intros n; elim n; auto; clear n.\n";
+ fprintf fmt " intros n Hrec ww ww_op; simpl GenBase.gen_to_Z; unfold zn2z_to_Z.\n";
+ fprintf fmt " rewrite <- Hrec; auto.\n";
+ fprintf fmt " unfold GenBase.gen_wB; rewrite <- digits_gend; auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+
+
+ fprintf fmt " Theorem digits_nmake:forall n ww (w_op: znz_op ww), \n";
+ fprintf fmt " znz_digits (nmake_op _ w_op (S n)) = \n";
+ fprintf fmt " xO (znz_digits (nmake_op _ w_op n)).\n";
+ fprintf fmt " Proof.\n";
+ fprintf fmt " auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+
+
+ fprintf fmt " Theorem znz_nmake_op: forall ww ww_op n xh xl,\n";
+ fprintf fmt " znz_to_Z (nmake_op ww ww_op (S n)) (WW xh xl) =\n";
+ fprintf fmt " znz_to_Z (nmake_op ww ww_op n) xh *\n";
+ fprintf fmt " base (znz_digits (nmake_op ww ww_op n)) +\n";
+ fprintf fmt " znz_to_Z (nmake_op ww ww_op n) xl.\n";
+ fprintf fmt " Proof.\n";
+ fprintf fmt " auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+
+ fprintf fmt " Theorem make_op_S: forall n,\n";
+ fprintf fmt " make_op (S n) = mk_zn2z_op_karatsuba (make_op n).\n";
+ fprintf fmt " intro n; pattern n; apply lt_wf_ind; clear n.\n";
+ fprintf fmt " intros n; case n; clear n.\n";
+ fprintf fmt " intros _; unfold make_op, make_op_aux, w%i_op; apply refl_equal.\n" (size + 2);
+ fprintf fmt " intros n; case n; clear n.\n";
+ fprintf fmt " intros _; unfold make_op, make_op_aux, w%i_op; apply refl_equal.\n" (size + 3);
+ fprintf fmt " intros n; case n; clear n.\n";
+ fprintf fmt " intros _; unfold make_op, make_op_aux, w%i_op, w%i_op; apply refl_equal.\n" (size + 3) (size + 2);
+ fprintf fmt " intros n Hrec.\n";
+ fprintf fmt " change (make_op (S (S (S (S n))))) with\n";
+ fprintf fmt " (mk_zn2z_op_karatsuba (mk_zn2z_op_karatsuba (mk_zn2z_op_karatsuba (make_op (S n))))).\n";
+ fprintf fmt " change (make_op (S (S (S n)))) with\n";
+ fprintf fmt " (mk_zn2z_op_karatsuba (mk_zn2z_op_karatsuba (mk_zn2z_op_karatsuba (make_op n)))).\n";
+ fprintf fmt " rewrite Hrec; auto with arith.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt " \n";
+
+
+ for i = 1 to size + 2 do
+ fprintf fmt " Let znz_to_Z_%i: forall x y,\n" i;
+ fprintf fmt " znz_to_Z w%i_op (WW x y) = \n" i;
+ fprintf fmt " znz_to_Z w%i_op x * base (znz_digits w%i_op) + znz_to_Z w%i_op y.\n" (i-1) (i-1) (i-1);
+ fprintf fmt " Proof.\n";
+ fprintf fmt " auto.\n";
+ fprintf fmt " Qed. \n";
+ fprintf fmt "\n";
done;
- fprintf fmt " | %s%i wx =>\n" c size;
- fprintf fmt " match w%i_succ_c wx with\n" size;
- fprintf fmt " | C0 r => %s%i r\n" c size;
- fprintf fmt " | C1 r => %sn 0 (WW one%i r)\n" c size ;
- fprintf fmt " end\n";
- fprintf fmt " | %sn n wx =>\n" c;
- fprintf fmt " let op := make_op n in\n";
- fprintf fmt " match op.(znz_succ_c) wx with\n";
- fprintf fmt " | C0 r => %sn n r\n" c;
- fprintf fmt " | C1 r => %sn (S n) (WW op.(znz_1) r)\n" c;
- fprintf fmt " end\n";
- fprintf fmt " end.\n";
+
+ fprintf fmt " Let znz_to_Z_n: forall n x y,\n";
+ fprintf fmt " znz_to_Z (make_op (S n)) (WW x y) = \n";
+ fprintf fmt " znz_to_Z (make_op n) x * base (znz_digits (make_op n)) + znz_to_Z (make_op n) y.\n";
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros n x y; rewrite make_op_S; auto.\n";
+ fprintf fmt " Qed. \n";
fprintf fmt "\n";
+ end;
- for i = 1 to size do
- fprintf fmt " Definition extend%i :=\n" i;
- fprintf fmt " Eval lazy beta zeta iota delta [extend]in extend %i.\n" i
+ if gen_proof then
+ begin
+ fprintf fmt " Let w0_spec: znz_spec w0_op := W0.w_spec.\n";
+ for i = 1 to 3 do
+ fprintf fmt " Let w%i_spec: znz_spec w%i_op := mk_znz2_spec w%i_spec.\n" i i (i-1)
+ done;
+ for i = 4 to size + 3 do
+ fprintf fmt " Let w%i_spec : znz_spec w%i_op := mk_znz2_karatsuba_spec w%i_spec.\n" i i (i-1)
done;
fprintf fmt "\n";
+
+ fprintf fmt " Let wn_spec: forall n, znz_spec (make_op n).\n";
+ fprintf fmt " intros n; elim n; clear n.\n";
+ fprintf fmt " exact w%i_spec.\n" (size + 1);
+ fprintf fmt " intros n Hrec; rewrite make_op_S.\n";
+ fprintf fmt " exact (mk_znz2_karatsuba_spec Hrec).\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt " \n";
+ end;
+
+ for i = 0 to size do
+ fprintf fmt " Definition w%i_eq0 := w%i_op.(znz_eq0).\n" i i;
+ fprintf fmt " Let spec_w%i_eq0: forall x, if w%i_eq0 x then [%s%i x] = 0 else True.\n" i i c i;
+ if gen_proof then
+ begin
+ fprintf fmt " intros x; unfold w%i_eq0, to_Z; generalize (spec_eq0 w%i_spec x);\n" i i;
+ fprintf fmt " case znz_eq0; auto.\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+ done;
+ fprintf fmt "\n";
+
+ if gen_proof then
+ begin
for i = 0 to size do
- fprintf fmt " Definition w%i_eq0 := w%i_op.(znz_eq0).\n" i i
+ fprintf fmt " Theorem digits_w%i: znz_digits w%i_op = znz_digits (nmake_op _ w0_op %i).\n" i i i;
+ if i == 0 then
+ fprintf fmt " auto.\n"
+ else
+ fprintf fmt " rewrite digits_nmake; rewrite <- digits_w%i; auto.\n" (i - 1);
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+
+ fprintf fmt " Let spec_gen_eval%in: forall n, eval%in n = GenBase.gen_to_Z (znz_digits w%i_op) (znz_to_Z w%i_op) n.\n" i i i i;
+ if gen_proof then
+ begin
+ fprintf fmt " intros n; exact (nmake_gen n w%i w%i_op).\n" i i;
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+ done;
+
+ for i = 0 to size do
+ for j = 0 to (size - i) do
+ fprintf fmt " Theorem digits_w%in%i: znz_digits w%i_op = znz_digits (nmake_op _ w%i_op %i).\n" i j (i + j) i j;
+ if j == 0 then
+ if i == 0 then
+ fprintf fmt " auto.\n"
+ else
+ begin
+ fprintf fmt " apply trans_equal with (xO (znz_digits w%i_op)).\n" (i + j -1);
+ fprintf fmt " auto.\n";
+ fprintf fmt " unfold nmake_op; auto.\n";
+ end
+ else
+ begin
+ fprintf fmt " apply trans_equal with (xO (znz_digits w%i_op)).\n" (i + j -1);
+ fprintf fmt " auto.\n";
+ fprintf fmt " rewrite digits_nmake.\n";
+ fprintf fmt " rewrite digits_w%in%i.\n" i (j - 1);
+ fprintf fmt " auto.\n";
+ end;
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+ fprintf fmt " Let spec_eval%in%i: forall x, [%s%i x] = eval%in %i x.\n" i j c (i + j) i j;
+ if gen_proof then
+ begin
+ if j == 0 then
+ fprintf fmt " intros x; rewrite spec_gen_eval%in; unfold GenBase.gen_to_Z, to_Z; auto.\n" i
+ else
+ begin
+ fprintf fmt " intros x; case x.\n";
+ fprintf fmt " auto.\n";
+ fprintf fmt " intros xh xl; unfold to_Z; rewrite znz_to_Z_%i.\n" (i + j);
+ fprintf fmt " rewrite digits_w%in%i.\n" i (j - 1);
+ fprintf fmt " generalize (spec_eval%in%i); unfold to_Z; intros HH; repeat rewrite HH.\n" i (j - 1);
+ fprintf fmt " unfold eval%in, nmake_op%i.\n" i i;
+ fprintf fmt " rewrite (znz_nmake_op _ w%i_op %i); auto.\n" i (j - 1);
+
+ end;
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ if i + j <> size then
+ begin
+ fprintf fmt " Let spec_extend%in%i: forall x, [%s%i x] = [%s%i (extend%i %i x)].\n" i (i + j + 1) c i c (i + j + 1) i j;
+ if j == 0 then
+ begin
+ fprintf fmt " intros x; change (extend%i 0 x) with (WW (znz_0 w%i_op) x).\n" i (i + j);
+ fprintf fmt " unfold to_Z; rewrite znz_to_Z_%i.\n" (i + j + 1);
+ fprintf fmt " rewrite (spec_0 w%i_spec); auto.\n" (i + j);
+
+ end
+ else
+ begin
+ fprintf fmt " intros x; change (extend%i %i x) with (WW (znz_0 w%i_op) (extend%i %i x)).\n" i j (i + j) i (j - 1);
+ fprintf fmt " unfold to_Z; rewrite znz_to_Z_%i.\n" (i + j + 1);
+ fprintf fmt " rewrite (spec_0 w%i_spec).\n" (i + j);
+ fprintf fmt " generalize (spec_extend%in%i x); unfold to_Z.\n" i (i + j);
+ fprintf fmt " intros HH; rewrite <- HH; auto.\n";
+
+ end;
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+ end;
+ done;
+
+ fprintf fmt " Theorem digits_w%in%i: znz_digits w%i_op = znz_digits (nmake_op _ w%i_op %i).\n" i (size - i + 1) (size + 1) i (size - i + 1);
+ fprintf fmt " apply trans_equal with (xO (znz_digits w%i_op)).\n" size;
+ fprintf fmt " auto.\n";
+ fprintf fmt " rewrite digits_nmake.\n";
+ fprintf fmt " rewrite digits_w%in%i.\n" i (size - i);
+ fprintf fmt " auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+
+ fprintf fmt " Let spec_eval%in%i: forall x, [%sn 0 x] = eval%in %i x.\n" i (size - i + 1) c i (size - i + 1);
+ fprintf fmt " intros x; case x.\n";
+ fprintf fmt " auto.\n";
+ fprintf fmt " intros xh xl; unfold to_Z; rewrite znz_to_Z_%i.\n" (size + 1);
+ fprintf fmt " rewrite digits_w%in%i.\n" i (size - i);
+ fprintf fmt " generalize (spec_eval%in%i); unfold to_Z; intros HH; repeat rewrite HH.\n" i (size - i);
+ fprintf fmt " unfold eval%in, nmake_op%i.\n" i i;
+ fprintf fmt " rewrite (znz_nmake_op _ w%i_op %i); auto.\n" i (size - i);
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+
+ fprintf fmt " Let spec_eval%in%i: forall x, [%sn 1 x] = eval%in %i x.\n" i (size - i + 2) c i (size - i + 2);
+ fprintf fmt " intros x; case x.\n";
+ fprintf fmt " auto.\n";
+ fprintf fmt " intros xh xl; unfold to_Z; rewrite znz_to_Z_%i.\n" (size + 2);
+ fprintf fmt " rewrite digits_w%in%i.\n" i (size + 1 - i);
+ fprintf fmt " generalize (spec_eval%in%i); unfold to_Z; change (make_op 0) with (w%i_op); intros HH; repeat rewrite HH.\n" i (size + 1 - i) (size + 1);
+ fprintf fmt " unfold eval%in, nmake_op%i.\n" i i;
+ fprintf fmt " rewrite (znz_nmake_op _ w%i_op %i); auto.\n" i (size + 1 - i);
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+ done;
+
+ fprintf fmt " Let digits_w%in: forall n,\n" size;
+ fprintf fmt " znz_digits (make_op n) = znz_digits (nmake_op _ w%i_op (S n)).\n" size;
+ fprintf fmt " intros n; elim n; clear n.\n";
+ fprintf fmt " change (znz_digits (make_op 0)) with (xO (znz_digits w%i_op)).\n" size;
+ fprintf fmt " rewrite nmake_op_S; apply sym_equal; auto.\n";
+ fprintf fmt " intros n Hrec.\n";
+ fprintf fmt " replace (znz_digits (make_op (S n))) with (xO (znz_digits (make_op n))).\n";
+ fprintf fmt " rewrite Hrec.\n";
+ fprintf fmt " rewrite nmake_op_S; apply sym_equal; auto.\n";
+ fprintf fmt " rewrite make_op_S; apply sym_equal; auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+
+ fprintf fmt " Let spec_eval%in: forall n x, [%sn n x] = eval%in (S n) x.\n" size c size;
+ fprintf fmt " intros n; elim n; clear n.\n";
+ fprintf fmt " exact spec_eval%in1.\n" size;
+ fprintf fmt " intros n Hrec x; case x; clear x.\n";
+ fprintf fmt " unfold to_Z, eval%in, nmake_op%i.\n" size size;
+ fprintf fmt " rewrite make_op_S; rewrite nmake_op_S; auto.\n";
+ fprintf fmt " intros xh xl.\n";
+ fprintf fmt " unfold to_Z in Hrec |- *.\n";
+ fprintf fmt " rewrite znz_to_Z_n.\n";
+ fprintf fmt " rewrite digits_w%in.\n" size;
+ fprintf fmt " repeat rewrite Hrec.\n";
+ fprintf fmt " unfold eval%in, nmake_op%i.\n" size size;
+ fprintf fmt " apply sym_equal; rewrite nmake_op_S; auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+
+ fprintf fmt " Let spec_extend%in: forall n x, [%s%i x] = [%sn n (extend%i n x)].\n" size c size c size ;
+ fprintf fmt " intros n; elim n; clear n.\n";
+ fprintf fmt " intros x; change (extend%i 0 x) with (WW (znz_0 w%i_op) x).\n" size size;
+ fprintf fmt " unfold to_Z.\n";
+ fprintf fmt " change (make_op 0) with w%i_op.\n" (size + 1);
+ fprintf fmt " rewrite znz_to_Z_%i; rewrite (spec_0 w%i_spec); auto.\n" (size + 1) size;
+ fprintf fmt " intros n Hrec x.\n";
+ fprintf fmt " change (extend%i (S n) x) with (WW W0 (extend%i n x)).\n" size size;
+ fprintf fmt " unfold to_Z in Hrec |- *; rewrite znz_to_Z_n; auto.\n";
+ fprintf fmt " rewrite <- Hrec.\n";
+ fprintf fmt " replace (znz_to_Z (make_op n) W0) with 0; auto.\n";
+ fprintf fmt " case n; auto; intros; rewrite make_op_S; auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+ end;
+
+
+
+ fprintf fmt " Theorem to_Z_pos: forall x, 0 <= [x].\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros x; case x; clear x.\n";
+ for i = 0 to size do
+ fprintf fmt " intros x; case (spec_to_Z w%i_spec x); auto.\n" i;
done;
+ fprintf fmt " intros n x; case (spec_to_Z (wn_spec n) x); auto.\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
fprintf fmt "\n";
-
+
+
+ if gen_proof then
+ begin
+ fprintf fmt " Let spec_extendn_0: forall n wx, [%sn n (extend n _ wx)] = [%sn 0 wx].\n" c c;
+ fprintf fmt " intros n; elim n; auto.\n";
+ fprintf fmt " intros n1 Hrec wx; simpl extend; rewrite <- Hrec; auto.\n";
+ fprintf fmt " unfold to_Z.\n";
+ fprintf fmt " case n1; auto; intros n2; repeat rewrite make_op_S; auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt " Hint Rewrite spec_extendn_0: extr.\n";
+ fprintf fmt "\n";
+ fprintf fmt " Let spec_extendn0_0: forall n wx, [%sn (S n) (WW W0 wx)] = [%sn n wx].\n" c c;
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros n x; unfold to_Z.\n";
+ fprintf fmt " rewrite znz_to_Z_n.\n";
+ fprintf fmt " rewrite <- (Zplus_0_l (znz_to_Z (make_op n) x)).\n";
+ fprintf fmt " apply (f_equal2 Zplus); auto.\n";
+ fprintf fmt " case n; auto.\n";
+ fprintf fmt " intros n1; rewrite make_op_S; auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt " Hint Rewrite spec_extendn_0: extr.\n";
+ fprintf fmt "\n";
+ fprintf fmt " Let spec_extend_tr: forall m n (w: word _ (S n)),\n";
+ fprintf fmt " [%sn (m + n) (extend_tr w m)] = [%sn n w].\n" c c;
+ fprintf fmt " Proof.\n";
+ fprintf fmt " induction m; auto.\n";
+ fprintf fmt " intros n x; simpl extend_tr.\n";
+ fprintf fmt " simpl plus; rewrite spec_extendn0_0; auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt " Hint Rewrite spec_extend_tr: extr.\n";
+ fprintf fmt "\n";
+ fprintf fmt " Let spec_cast_l: forall n m x1,\n";
+ fprintf fmt " [%sn (Max.max n m)\n" c;
+ fprintf fmt " (castm (diff_r n m) (extend_tr x1 (snd (diff n m))))] =\n";
+ fprintf fmt " [%sn n x1].\n" c;
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros n m x1; case (diff_r n m); simpl castm.\n";
+ fprintf fmt " rewrite spec_extend_tr; auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt " Hint Rewrite spec_cast_l: extr.\n";
+ fprintf fmt "\n";
+ fprintf fmt " Let spec_cast_r: forall n m x1,\n";
+ fprintf fmt " [%sn (Max.max n m)\n" c;
+ fprintf fmt " (castm (diff_l n m) (extend_tr x1 (fst (diff n m))))] =\n";
+ fprintf fmt " [%sn m x1].\n" c;
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros n m x1; case (diff_l n m); simpl castm.\n";
+ fprintf fmt " rewrite spec_extend_tr; auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt " Hint Rewrite spec_cast_r: extr.\n";
+ fprintf fmt "\n";
+ end;
+
+
+ fprintf fmt " Section LevelAndIter.\n";
+ fprintf fmt "\n";
+ fprintf fmt " Variable res: Set.\n";
+ fprintf fmt " Variable xxx: res.\n";
+ fprintf fmt " Variable P: Z -> Z -> res -> Prop.\n";
+ fprintf fmt " (* Abstraction function for each level *)\n";
for i = 0 to size do
- fprintf fmt " Definition w%i_0W := w%i_op.(znz_0W).\n" i i
+ fprintf fmt " Variable f%i: w%i -> w%i -> res.\n" i i i;
+ fprintf fmt " Variable f%in: forall n, w%i -> word w%i (S n) -> res.\n" i i i;
+ fprintf fmt " Variable fn%i: forall n, word w%i (S n) -> w%i -> res.\n" i i i;
+ if gen_proof then
+ begin
+ fprintf fmt " Variable Pf%i: forall x y, P [%s%i x] [%s%i y] (f%i x y).\n" i c i c i i;
+ if i == size then
+ begin
+ fprintf fmt " Variable Pf%in: forall n x y, P [%s%i x] (eval%in (S n) y) (f%in n x y).\n" i c i i i;
+ fprintf fmt " Variable Pfn%i: forall n x y, P (eval%in (S n) x) [%s%i y] (fn%i n x y).\n" i i c i i;
+ end
+ else
+ begin
+
+ fprintf fmt " Variable Pf%in: forall n x y, Z_of_nat n <= %i -> P [%s%i x] (eval%in (S n) y) (f%in n x y).\n" i (size - i) c i i i;
+ fprintf fmt " Variable Pfn%i: forall n x y, Z_of_nat n <= %i -> P (eval%in (S n) x) [%s%i y] (fn%i n x y).\n" i (size - i) i c i i;
+ end;
+ end;
+ fprintf fmt "\n";
+ done;
+ fprintf fmt " Variable fnn: forall n, word w%i (S n) -> word w%i (S n) -> res.\n" size size;
+ if gen_proof then
+ fprintf fmt " Variable Pfnn: forall n x y, P [%sn n x] [%sn n y] (fnn n x y).\n" c c;
+ fprintf fmt " Variable fnm: forall n m, word w%i (S n) -> word w%i (S m) -> res.\n" size size;
+ if gen_proof then
+ fprintf fmt " Variable Pfnm: forall n m x y, P [%sn n x] [%sn m y] (fnm n m x y).\n" c c;
+ fprintf fmt "\n";
+ fprintf fmt " (* Special zero functions *)\n";
+ fprintf fmt " Variable f0t: t_ -> res.\n";
+ if gen_proof then
+ fprintf fmt " Variable Pf0t: forall x, P 0 [x] (f0t x).\n";
+ fprintf fmt " Variable ft0: t_ -> res.\n";
+ if gen_proof then
+ fprintf fmt " Variable Pft0: forall x, P [x] 0 (ft0 x).\n";
+ fprintf fmt "\n";
+
+
+ fprintf fmt " (* We level the two arguments before applying *)\n";
+ fprintf fmt " (* the functions at each leval *)\n";
+ fprintf fmt " Definition same_level (x y: t_): res :=\n";
+ fprintf fmt " Eval lazy zeta beta iota delta [";
+ for i = 0 to size do
+ fprintf fmt "extend%i " i;
done;
fprintf fmt "\n";
- fprintf fmt " Definition w0_WW := w0_op.(znz_WW).\n";
+ fprintf fmt " GenBase.extend GenBase.extend_aux\n";
+ fprintf fmt " ] in\n";
+ fprintf fmt " match x, y with\n";
+ for i = 0 to size do
+ for j = 0 to i - 1 do
+ fprintf fmt " | %s%i wx, %s%i wy => f%i wx (extend%i %i wy)\n" c i c j i j (i - j -1);
+ done;
+ fprintf fmt " | %s%i wx, %s%i wy => f%i wx wy\n" c i c i i;
+ for j = i + 1 to size do
+ fprintf fmt " | %s%i wx, %s%i wy => f%i (extend%i %i wx) wy\n" c i c j j i (j - i - 1);
+ done;
+ if i == size then
+ fprintf fmt " | %s%i wx, %sn m wy => fnn m (extend%i m wx) wy\n" c size c size
+ else
+ fprintf fmt " | %s%i wx, %sn m wy => fnn m (extend%i m (extend%i %i wx)) wy\n" c i c size i (size - i - 1);
+ done;
+ for i = 0 to size do
+ if i == size then
+ fprintf fmt " | %sn n wx, %s%i wy => fnn n wx (extend%i n wy)\n" c c size size
+ else
+ fprintf fmt " | %sn n wx, %s%i wy => fnn n wx (extend%i n (extend%i %i wy))\n" c c i size i (size - i - 1);
+ done;
+ fprintf fmt " | %sn n wx, Nn m wy =>\n" c;
+ fprintf fmt " let mn := Max.max n m in\n";
+ fprintf fmt " let d := diff n m in\n";
+ fprintf fmt " fnn mn\n";
+ fprintf fmt " (castm (diff_r n m) (extend_tr wx (snd d)))\n";
+ fprintf fmt " (castm (diff_l n m) (extend_tr wy (fst d)))\n";
+ fprintf fmt " end.\n";
+ fprintf fmt "\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Lemma spec_same_level: forall x y, P [x] [y] (same_level x y).\n";
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros x; case x; clear x; unfold same_level.\n";
+ for i = 0 to size do
+ fprintf fmt " intros x y; case y; clear y.\n";
+ for j = 0 to i - 1 do
+ fprintf fmt " intros y; rewrite spec_extend%in%i; apply Pf%i.\n" j i i;
+ done;
+ fprintf fmt " intros y; apply Pf%i.\n" i;
+ for j = i + 1 to size do
+ fprintf fmt " intros y; rewrite spec_extend%in%i; apply Pf%i.\n" i j j;
+ done;
+ if i == size then
+ fprintf fmt " intros m y; rewrite (spec_extend%in m); apply Pfnn.\n" size
+ else
+ fprintf fmt " intros m y; rewrite spec_extend%in%i; rewrite (spec_extend%in m); apply Pfnn.\n" i size size;
+ done;
+ fprintf fmt " intros n x y; case y; clear y.\n";
+ for i = 0 to size do
+ if i == size then
+ fprintf fmt " intros y; rewrite (spec_extend%in n); apply Pfnn.\n" size
+ else
+ fprintf fmt " intros y; rewrite spec_extend%in%i; rewrite (spec_extend%in n); apply Pfnn.\n" i size size;
+ done;
+ fprintf fmt " intros m y; rewrite <- (spec_cast_l n m x); \n";
+ fprintf fmt " rewrite <- (spec_cast_r n m y); apply Pfnn.\n";
+ fprintf fmt " Qed.\n";
fprintf fmt "\n";
+ end;
- (* Addition *)
+ fprintf fmt " (* We level the two arguments before applying *)\n";
+ fprintf fmt " (* the functions at each level (special zero case) *)\n";
+ fprintf fmt " Definition same_level0 (x y: t_): res :=\n";
+ fprintf fmt " Eval lazy zeta beta iota delta [";
for i = 0 to size do
- fprintf fmt " Definition w%i_add_c := w%i_op.(znz_add_c).\n" i i
+ fprintf fmt "extend%i " i;
done;
fprintf fmt "\n";
-(*
- fprintf fmt " Definition add_c_1_0 x y :=\n";
+ fprintf fmt " GenBase.extend GenBase.extend_aux\n";
+ fprintf fmt " ] in\n";
fprintf fmt " match x with\n";
- fprintf fmt " | W0 => C0 (w0_0W y)\n";
- fprintf fmt " | WW xh xl =>
- fprintf fmt " match w1_add_c xl y with\n";
- fprintf fmt " | C0 rl => C0 (WW xh rl)\n";
- fprintf fmt " | C1 rl =>\n";
- fprintf fmt " match w1_succ_c xh with\n";
- fprintf fmt " | C0 rh => C0 (WW rh rl)\n";
- fprintf fmt " | C1 rh => C1 (w0_WW rh rl)\n";
- fprintf fmt " end\n";
+ for i = 0 to size do
+ fprintf fmt " | %s%i wx =>\n" c i;
+ if (i == 0) then
+ fprintf fmt " if w0_eq0 wx then f0t y else\n";
+ fprintf fmt " match y with\n";
+ for j = 0 to i - 1 do
+ fprintf fmt " | %s%i wy =>\n" c j;
+ if j == 0 then
+ fprintf fmt " if w0_eq0 wy then ft0 x else\n";
+ fprintf fmt " f%i wx (extend%i %i wy)\n" i j (i - j -1);
+ done;
+ fprintf fmt " | %s%i wy => f%i wx wy\n" c i i;
+ for j = i + 1 to size do
+ fprintf fmt " | %s%i wy => f%i (extend%i %i wx) wy\n" c j j i (j - i - 1);
+ done;
+ if i == size then
+ fprintf fmt " | %sn m wy => fnn m (extend%i m wx) wy\n" c size
+ else
+ fprintf fmt " | %sn m wy => fnn m (extend%i m (extend%i %i wx)) wy\n" c size i (size - i - 1);
+ fprintf fmt" end\n";
+ done;
+ fprintf fmt " | %sn n wx =>\n" c;
+ fprintf fmt " match y with\n";
+ for i = 0 to size do
+ fprintf fmt " | %s%i wy =>\n" c i;
+ if i == 0 then
+ fprintf fmt " if w0_eq0 wy then ft0 x else\n";
+ if i == size then
+ fprintf fmt " fnn n wx (extend%i n wy)\n" size
+ else
+ fprintf fmt " fnn n wx (extend%i n (extend%i %i wy))\n" size i (size - i - 1);
+ done;
+ fprintf fmt " | %sn m wy =>\n" c;
+ fprintf fmt " let mn := Max.max n m in\n";
+ fprintf fmt " let d := diff n m in\n";
+ fprintf fmt " fnn mn\n";
+ fprintf fmt " (castm (diff_r n m) (extend_tr wx (snd d)))\n";
+ fprintf fmt " (castm (diff_l n m) (extend_tr wy (fst d)))\n";
fprintf fmt " end\n";
fprintf fmt " end.\n";
fprintf fmt "\n";
- for i = 1 to size do
- fprintf fmt " Definition add_c_n_%i :=\n" i;
- fprintf fmt " add_c_smn1 w%i
-*)
-
- for i = 0 to size do
- fprintf fmt " Definition w%i_add x y :=\n" i;
- fprintf fmt " match w%i_add_c x y with\n" i;
- fprintf fmt " | C0 r => %s%i r\n" c i;
- fprintf fmt " | C1 r => ";
- if i < size then fprintf fmt "%s%i (WW one%i r)\n" c (i+1) i
- else fprintf fmt "%sn 0 (WW one%i r)\n" c size;
- fprintf fmt " end.\n"
+ if gen_proof then
+ begin
+ fprintf fmt " Lemma spec_same_level0: forall x y, P [x] [y] (same_level0 x y).\n";
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros x; case x; clear x; unfold same_level0.\n";
+ for i = 0 to size do
+ fprintf fmt " intros x.\n";
+ if i == 0 then
+ begin
+ fprintf fmt " generalize (spec_w0_eq0 x); case w0_eq0; intros H.\n";
+ fprintf fmt " intros y; rewrite H; apply Pf0t.\n";
+ fprintf fmt " clear H.\n";
+ end;
+ fprintf fmt " intros y; case y; clear y.\n";
+ for j = 0 to i - 1 do
+ fprintf fmt " intros y.\n";
+ if j == 0 then
+ begin
+ fprintf fmt " generalize (spec_w0_eq0 y); case w0_eq0; intros H.\n";
+ fprintf fmt " rewrite H; apply Pft0.\n";
+ fprintf fmt " clear H.\n";
+ end;
+ fprintf fmt " rewrite spec_extend%in%i; apply Pf%i.\n" j i i;
+ done;
+ fprintf fmt " intros y; apply Pf%i.\n" i;
+ for j = i + 1 to size do
+ fprintf fmt " intros y; rewrite spec_extend%in%i; apply Pf%i.\n" i j j;
+ done;
+ if i == size then
+ fprintf fmt " intros m y; rewrite (spec_extend%in m); apply Pfnn.\n" size
+ else
+ fprintf fmt " intros m y; rewrite spec_extend%in%i; rewrite (spec_extend%in m); apply Pfnn.\n" i size size;
done;
- fprintf fmt " Definition addn n (x y : word w%i (S n)) :=\n" size;
- fprintf fmt " let op := make_op n in\n";
- fprintf fmt " match op.(znz_add_c) x y with\n";
- fprintf fmt " | C0 r => %sn n r\n" c;
- fprintf fmt " | C1 r => %sn (S n) (WW op.(znz_1) r)" c;
- fprintf fmt " end.\n";
+ fprintf fmt " intros n x y; case y; clear y.\n";
+ for i = 0 to size do
+ fprintf fmt " intros y.\n";
+ if i = 0 then
+ begin
+ fprintf fmt " generalize (spec_w0_eq0 y); case w0_eq0; intros H.\n";
+ fprintf fmt " rewrite H; apply Pft0.\n";
+ fprintf fmt " clear H.\n";
+ end;
+ if i == size then
+ fprintf fmt " rewrite (spec_extend%in n); apply Pfnn.\n" size
+ else
+ fprintf fmt " rewrite spec_extend%in%i; rewrite (spec_extend%in n); apply Pfnn.\n" i size size;
+ done;
+ fprintf fmt " intros m y; rewrite <- (spec_cast_l n m x); \n";
+ fprintf fmt " rewrite <- (spec_cast_r n m y); apply Pfnn.\n";
+ fprintf fmt " Qed.\n";
fprintf fmt "\n";
+ end;
- fprintf fmt " Definition add x y :=\n";
+ fprintf fmt " (* We iter the smaller argument with the bigger *)\n";
+ fprintf fmt " Definition iter (x y: t_): res := \n";
+ fprintf fmt " Eval lazy zeta beta iota delta [";
+ for i = 0 to size do
+ fprintf fmt "extend%i " i;
+ done;
+ fprintf fmt "\n";
+ fprintf fmt " GenBase.extend GenBase.extend_aux\n";
+ fprintf fmt " ] in\n";
fprintf fmt " match x, y with\n";
- fprintf fmt " | %s0 wx, %s0 wy => w0_add wx wy \n" c c;
- for j = 1 to size do
- fprintf fmt " | %s0 wx, %s%i wy =>\n" c c j;
- fprintf fmt " if w0_eq0 wx then y else w%i_add " j;
- if j = 1 then fprintf fmt "(WW w_0 wx) wy\n"
- else fprintf fmt "(extend%i w0 (WW w_0 wx)) wy\n" (j-1)
- done;
- fprintf fmt " | %s0 wx, %sn n wy =>\n" c c;
- fprintf fmt " if w0_eq0 wx then y\n";
- fprintf fmt " else addn n (extend n w%i (extend%i w0 (WW w_0 wx))) wy\n"
- size size;
- for i = 1 to size do
- fprintf fmt " | %s%i wx, %s0 wy =>\n" c i c;
- fprintf fmt
- " if w0_eq0 wy then x else w%i_add wx " i;
- if i = 1 then fprintf fmt "(WW w_0 wy)\n"
- else fprintf fmt "(extend%i w0 (WW w_0 wy))\n" (i-1);
- for j = 1 to size do
- fprintf fmt " | %s%i wx, %s%i wy => " c i c j;
- if i < j then fprintf fmt "w%i_add (extend%i w%i wx) wy\n" j (j-i) (i-1)
- else if i = j then fprintf fmt "w%i_add wx wy\n" j
- else fprintf fmt "w%i_add wx (extend%i w%i wy)\n" i (i-j) (j-1)
+ for i = 0 to size do
+ for j = 0 to i - 1 do
+ fprintf fmt " | %s%i wx, %s%i wy => fn%i %i wx wy\n" c i c j j (i - j - 1);
done;
- fprintf fmt
- " | %s%i wx, %sn n wy => addn n (extend n w%i (extend%i w%i wx)) wy\n"
- c i c size (size-i+1) (i-1)
- done;
- fprintf fmt " | %sn n wx, %s0 wy =>\n" c c;
- fprintf fmt " if w0_eq0 wy then x\n";
- fprintf fmt " else addn n wx (extend n w%i (extend%i w0 (WW w_0 wy)))\n"
- size size;
- for j = 1 to size do
- fprintf fmt
- " | %sn n wx, %s%i wy => addn n wx (extend n w%i (extend%i w%i wy))\n"
- c c j size (size-j+1) (j-1);
- done;
- fprintf fmt " | %sn n wx, %sn m wy =>\n" c c;
- fprintf fmt " let mn := Max.max n m in\n";
- fprintf fmt " let d := diff n m in\n";
- fprintf fmt " addn mn\n";
- fprintf fmt " (castm (diff_r n m) (extend_tr wx (snd d)))\n";
- fprintf fmt " (castm (diff_l n m) (extend_tr wy (fst d)))\n";
+ fprintf fmt " | %s%i wx, %s%i wy => f%i wx wy\n" c i c i i;
+ for j = i + 1 to size do
+ fprintf fmt " | %s%i wx, %s%i wy => f%in %i wx wy\n" c i c j i (j - i - 1);
+ done;
+ if i == size then
+ fprintf fmt " | %s%i wx, %sn m wy => f%in m wx wy\n" c size c size
+ else
+ fprintf fmt " | %s%i wx, %sn m wy => f%in m (extend%i %i wx) wy\n" c i c size i (size - i - 1);
+ done;
+ for i = 0 to size do
+ if i == size then
+ fprintf fmt " | %sn n wx, %s%i wy => fn%i n wx wy\n" c c size size
+ else
+ fprintf fmt " | %sn n wx, %s%i wy => fn%i n wx (extend%i %i wy)\n" c c i size i (size - i - 1);
+ done;
+ fprintf fmt " | %sn n wx, %sn m wy => fnm n m wx wy\n" c c;
fprintf fmt " end.\n";
fprintf fmt "\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Ltac zg_tac := try\n";
+ fprintf fmt " (red; simpl Zcompare; auto;\n";
+ fprintf fmt " let t := fresh \"H\" in (intros t; discriminate H)).\n";
+ fprintf fmt " Lemma spec_iter: forall x y, P [x] [y] (iter x y).\n";
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros x; case x; clear x; unfold iter.\n";
+ for i = 0 to size do
+ fprintf fmt " intros x y; case y; clear y.\n";
+ for j = 0 to i - 1 do
+ fprintf fmt " intros y; rewrite spec_eval%in%i; apply (Pfn%i %i); zg_tac.\n" j (i - j) j (i - j - 1);
+ done;
+ fprintf fmt " intros y; apply Pf%i.\n" i;
+ for j = i + 1 to size do
+ fprintf fmt " intros y; rewrite spec_eval%in%i; apply (Pf%in %i); zg_tac.\n" i (j - i) i (j - i - 1);
+ done;
+ if i == size then
+ fprintf fmt " intros m y; rewrite spec_eval%in; apply Pf%in.\n" size size
+ else
+ fprintf fmt " intros m y; rewrite spec_extend%in%i; rewrite spec_eval%in; apply Pf%in.\n" i size size size;
+ done;
+ fprintf fmt " intros n x y; case y; clear y.\n";
+ for i = 0 to size do
+ if i == size then
+ fprintf fmt " intros y; rewrite spec_eval%in; apply Pfn%i.\n" size size
+ else
+ fprintf fmt " intros y; rewrite spec_extend%in%i; rewrite spec_eval%in; apply Pfn%i.\n" i size size size;
+ done;
+ fprintf fmt " intros m y; apply Pfnm.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+ end;
+
+
+ fprintf fmt " (* We iter the smaller argument with the bigger (zero case) *)\n";
+ fprintf fmt " Definition iter0 (x y: t_): res :=\n";
+ fprintf fmt " Eval lazy zeta beta iota delta [";
+ for i = 0 to size do
+ fprintf fmt "extend%i " i;
+ done;
+ fprintf fmt "\n";
+ fprintf fmt " GenBase.extend GenBase.extend_aux\n";
+ fprintf fmt " ] in\n";
+ fprintf fmt " match x with\n";
+ for i = 0 to size do
+ fprintf fmt " | %s%i wx =>\n" c i;
+ if (i == 0) then
+ fprintf fmt " if w0_eq0 wx then f0t y else\n";
+ fprintf fmt " match y with\n";
+ for j = 0 to i - 1 do
+ fprintf fmt " | %s%i wy =>\n" c j;
+ if j == 0 then
+ fprintf fmt " if w0_eq0 wy then ft0 x else\n";
+ fprintf fmt " fn%i %i wx wy\n" j (i - j - 1);
+ done;
+ fprintf fmt " | %s%i wy => f%i wx wy\n" c i i;
+ for j = i + 1 to size do
+ fprintf fmt " | %s%i wy => f%in %i wx wy\n" c j i (j - i - 1);
+ done;
+ if i == size then
+ fprintf fmt " | %sn m wy => f%in m wx wy\n" c size
+ else
+ fprintf fmt " | %sn m wy => f%in m (extend%i %i wx) wy\n" c size i (size - i - 1);
+ fprintf fmt " end\n";
+ done;
+ fprintf fmt " | %sn n wx =>\n" c;
+ fprintf fmt " match y with\n";
+ for i = 0 to size do
+ fprintf fmt " | %s%i wy =>\n" c i;
+ if i == 0 then
+ fprintf fmt " if w0_eq0 wy then ft0 x else\n";
+ if i == size then
+ fprintf fmt " fn%i n wx wy\n" size
+ else
+ fprintf fmt " fn%i n wx (extend%i %i wy)\n" size i (size - i - 1);
+ done;
+ fprintf fmt " | %sn m wy => fnm n m wx wy\n" c;
+ fprintf fmt " end\n";
+ fprintf fmt " end.\n";
+ fprintf fmt "\n";
+
+ if gen_proof then
+ begin
+ fprintf fmt " Lemma spec_iter0: forall x y, P [x] [y] (iter0 x y).\n";
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros x; case x; clear x; unfold iter0.\n";
+ for i = 0 to size do
+ fprintf fmt " intros x.\n";
+ if i == 0 then
+ begin
+ fprintf fmt " generalize (spec_w0_eq0 x); case w0_eq0; intros H.\n";
+ fprintf fmt " intros y; rewrite H; apply Pf0t.\n";
+ fprintf fmt " clear H.\n";
+ end;
+ fprintf fmt " intros y; case y; clear y.\n";
+ for j = 0 to i - 1 do
+ fprintf fmt " intros y.\n";
+ if j == 0 then
+ begin
+ fprintf fmt " generalize (spec_w0_eq0 y); case w0_eq0; intros H.\n";
+ fprintf fmt " rewrite H; apply Pft0.\n";
+ fprintf fmt " clear H.\n";
+ end;
+ fprintf fmt " rewrite spec_eval%in%i; apply (Pfn%i %i); zg_tac.\n" j (i - j) j (i - j - 1);
+ done;
+ fprintf fmt " intros y; apply Pf%i.\n" i;
+ for j = i + 1 to size do
+ fprintf fmt " intros y; rewrite spec_eval%in%i; apply (Pf%in %i); zg_tac.\n" i (j - i) i (j - i - 1);
+ done;
+ if i == size then
+ fprintf fmt " intros m y; rewrite spec_eval%in; apply Pf%in.\n" size size
+ else
+ fprintf fmt " intros m y; rewrite spec_extend%in%i; rewrite spec_eval%in; apply Pf%in.\n" i size size size;
+ done;
+ fprintf fmt " intros n x y; case y; clear y.\n";
+ for i = 0 to size do
+ fprintf fmt " intros y.\n";
+ if i = 0 then
+ begin
+ fprintf fmt " generalize (spec_w0_eq0 y); case w0_eq0; intros H.\n";
+ fprintf fmt " rewrite H; apply Pft0.\n";
+ fprintf fmt " clear H.\n";
+ end;
+ if i == size then
+ fprintf fmt " rewrite spec_eval%in; apply Pfn%i.\n" size size
+ else
+ fprintf fmt " rewrite spec_extend%in%i; rewrite spec_eval%in; apply Pfn%i.\n" i size size size;
+ done;
+ fprintf fmt " intros m y; apply Pfnm.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+ end;
+
+
+ fprintf fmt " End LevelAndIter.\n";
+ fprintf fmt "\n";
+
+
+ fprintf fmt " (***************************************************************)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (* Reduction *)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (***************************************************************)\n\n";
+
fprintf fmt " Definition reduce_0 (x:w) := %s0 x.\n" c;
fprintf fmt " Definition reduce_1 :=\n";
fprintf fmt " Eval lazy beta iota delta[reduce_n1] in\n";
@@ -286,7 +941,199 @@ let print_Make () =
fprintf fmt " reduce_n _ _ zero reduce_%i %sn n.\n" (size + 1) c;
fprintf fmt "\n";
- (* Predecessor *)
+ if gen_proof then
+ begin
+ fprintf fmt " Let spec_reduce_0: forall x, [reduce_0 x] = [%s0 x].\n" c;
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros x; unfold to_Z, reduce_0.\n";
+ fprintf fmt " auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt " \n";
+
+ for i = 1 to size + 1 do
+ if (i == size + 1) then
+ fprintf fmt " Let spec_reduce_%i: forall x, [reduce_%i x] = [%sn 0 x].\n" i i c
+ else
+ fprintf fmt " Let spec_reduce_%i: forall x, [reduce_%i x] = [%s%i x].\n" i i c i;
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros x; case x; unfold reduce_%i.\n" i;
+ fprintf fmt " exact (spec_0 w0_spec).\n";
+ fprintf fmt " intros x1 y1.\n";
+ fprintf fmt " generalize (spec_w%i_eq0 x1); \n" (i - 1);
+ fprintf fmt " case w%i_eq0; intros H1; auto.\n" (i - 1);
+ if i <> 1 then
+ fprintf fmt " rewrite spec_reduce_%i.\n" (i - 1);
+ fprintf fmt " unfold to_Z; rewrite znz_to_Z_%i.\n" i;
+ fprintf fmt " unfold to_Z in H1; rewrite H1; auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt " \n";
+ done;
+
+ fprintf fmt " Let spec_reduce_n: forall n x, [reduce_n n x] = [%sn n x].\n" c;
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros n; elim n; simpl reduce_n.\n";
+ fprintf fmt " intros x; rewrite <- spec_reduce_%i; auto.\n" (size + 1);
+ fprintf fmt " intros n1 Hrec x; case x.\n";
+ fprintf fmt " unfold to_Z; rewrite make_op_S; auto.\n";
+ fprintf fmt " exact (spec_0 w0_spec).\n";
+ fprintf fmt " intros x1 y1; case x1; auto.\n";
+ fprintf fmt " rewrite Hrec.\n";
+ fprintf fmt " rewrite spec_extendn0_0; auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt " \n";
+ end;
+
+ fprintf fmt " (***************************************************************)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (* Successor *)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (***************************************************************)\n\n";
+
+ for i = 0 to size do
+ fprintf fmt " Definition w%i_succ_c := w%i_op.(znz_succ_c).\n" i i
+ done;
+ fprintf fmt "\n";
+
+ for i = 0 to size do
+ fprintf fmt " Definition w%i_succ := w%i_op.(znz_succ).\n" i i
+ done;
+ fprintf fmt "\n";
+
+ fprintf fmt " Definition succ x :=\n";
+ fprintf fmt " match x with\n";
+ for i = 0 to size-1 do
+ fprintf fmt " | %s%i wx =>\n" c i;
+ fprintf fmt " match w%i_succ_c wx with\n" i;
+ fprintf fmt " | C0 r => %s%i r\n" c i;
+ fprintf fmt " | C1 r => %s%i (WW one%i r)\n" c (i+1) i;
+ fprintf fmt " end\n";
+ done;
+ fprintf fmt " | %s%i wx =>\n" c size;
+ fprintf fmt " match w%i_succ_c wx with\n" size;
+ fprintf fmt " | C0 r => %s%i r\n" c size;
+ fprintf fmt " | C1 r => %sn 0 (WW one%i r)\n" c size ;
+ fprintf fmt " end\n";
+ fprintf fmt " | %sn n wx =>\n" c;
+ fprintf fmt " let op := make_op n in\n";
+ fprintf fmt " match op.(znz_succ_c) wx with\n";
+ fprintf fmt " | C0 r => %sn n r\n" c;
+ fprintf fmt " | C1 r => %sn (S n) (WW op.(znz_1) r)\n" c;
+ fprintf fmt " end\n";
+ fprintf fmt " end.\n";
+ fprintf fmt "\n";
+
+ fprintf fmt " Theorem succ_spec: forall n, [succ n] = [n] + 1.\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros n; case n; unfold succ, to_Z.\n";
+ for i = 0 to size do
+ fprintf fmt " intros n1; generalize (spec_succ_c w%i_spec n1);\n" i;
+ fprintf fmt " unfold succ, to_Z, w%i_succ_c; case znz_succ_c; auto.\n" i;
+ fprintf fmt " intros ww H; rewrite <- H.\n";
+ fprintf fmt " (rewrite znz_to_Z_%i; unfold interp_carry;\n" (i + 1);
+ fprintf fmt " apply f_equal2 with (f := Zplus); auto;\n";
+ fprintf fmt " apply f_equal2 with (f := Zmult); auto;\n";
+ fprintf fmt " exact (spec_1 w%i_spec)).\n" i;
+ done;
+ fprintf fmt " intros k n1; generalize (spec_succ_c (wn_spec k) n1).\n";
+ fprintf fmt " unfold succ, to_Z; case znz_succ_c; auto.\n";
+ fprintf fmt " intros ww H; rewrite <- H.\n";
+ fprintf fmt " (rewrite (znz_to_Z_n k); unfold interp_carry;\n";
+ fprintf fmt " apply f_equal2 with (f := Zplus); auto;\n";
+ fprintf fmt " apply f_equal2 with (f := Zmult); auto;\n";
+ fprintf fmt " exact (spec_1 (wn_spec k))).\n";
+ fprintf fmt " Qed.\n";
+ end
+ else fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+
+
+ fprintf fmt " (***************************************************************)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (* Adddition *)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (***************************************************************)\n\n";
+ for i = 0 to size do
+ fprintf fmt " Definition w%i_add_c := znz_add_c w%i_op.\n" i i;
+ fprintf fmt " Definition w%i_add x y :=\n" i;
+ fprintf fmt " match w%i_add_c x y with\n" i;
+ fprintf fmt " | C0 r => %s%i r\n" c i;
+ if i == size then
+ fprintf fmt " | C1 r => %sn 0 (WW one%i r)\n" c size
+ else
+ fprintf fmt " | C1 r => %s%i (WW one%i r)\n" c (i + 1) i;
+ fprintf fmt " end.\n";
+ fprintf fmt "\n";
+ done ;
+ fprintf fmt " Definition addn n (x y : word w%i (S n)) :=\n" size;
+ fprintf fmt " let op := make_op n in\n";
+ fprintf fmt " match op.(znz_add_c) x y with\n";
+ fprintf fmt " | C0 r => %sn n r\n" c;
+ fprintf fmt " | C1 r => %sn (S n) (WW op.(znz_1) r) end.\n" c;
+ fprintf fmt "\n";
+
+
+ if gen_proof then
+ begin
+ for i = 0 to size do
+ fprintf fmt " Let spec_w%i_add: forall x y, [w%i_add x y] = [%s%i x] + [%s%i y].\n" i i c i c i;
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros n m; unfold to_Z, w%i_add, w%i_add_c.\n" i i;
+ fprintf fmt " generalize (spec_add_c w%i_spec n m); case znz_add_c; auto.\n" i;
+ fprintf fmt " intros ww H; rewrite <- H.\n";
+ fprintf fmt " rewrite znz_to_Z_%i; unfold interp_carry;\n" (i + 1);
+ fprintf fmt " apply f_equal2 with (f := Zplus); auto;\n";
+ fprintf fmt " apply f_equal2 with (f := Zmult); auto;\n";
+ fprintf fmt " exact (spec_1 w%i_spec).\n" i;
+ fprintf fmt " Qed.\n";
+ fprintf fmt " Hint Rewrite spec_w%i_add: addr.\n" i;
+ fprintf fmt "\n";
+ done;
+ fprintf fmt " Let spec_wn_add: forall n x y, [addn n x y] = [%sn n x] + [%sn n y].\n" c c;
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros k n m; unfold to_Z, addn.\n";
+ fprintf fmt " generalize (spec_add_c (wn_spec k) n m); case znz_add_c; auto.\n";
+ fprintf fmt " intros ww H; rewrite <- H.\n";
+ fprintf fmt " rewrite (znz_to_Z_n k); unfold interp_carry;\n";
+ fprintf fmt " apply f_equal2 with (f := Zplus); auto;\n";
+ fprintf fmt " apply f_equal2 with (f := Zmult); auto;\n";
+ fprintf fmt " exact (spec_1 (wn_spec k)).\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt " Hint Rewrite spec_wn_add: addr.\n";
+ end;
+
+ fprintf fmt " Definition add := Eval lazy beta delta [same_level] in\n";
+ fprintf fmt " (same_level t_ ";
+ for i = 0 to size do
+ fprintf fmt "w%i_add " i;
+ done;
+ fprintf fmt "addn).\n";
+ fprintf fmt "\n";
+
+ fprintf fmt " Theorem spec_add: forall x y, [add x y] = [x] + [y].\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " unfold add.\n";
+ fprintf fmt " generalize (spec_same_level t_ (fun x y res => [res] = x + y)).\n";
+ fprintf fmt " unfold same_level; intros HH; apply HH; clear HH.\n";
+ for i = 0 to size do
+ fprintf fmt " exact spec_w%i_add.\n" i;
+ done;
+ fprintf fmt " exact spec_wn_add.\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+
+ fprintf fmt " (***************************************************************)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (* Predecessor *)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (***************************************************************)\n\n";
+
for i = 0 to size do
fprintf fmt " Definition w%i_pred_c := w%i_op.(znz_pred_c).\n" i i
done;
@@ -310,8 +1157,62 @@ let print_Make () =
fprintf fmt " end.\n";
fprintf fmt "\n";
- (* Substraction *)
- fprintf fmt "\n";
+ fprintf fmt " Let spec_pred: forall x, 0 < [x] -> [pred x] = [x] - 1.\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros x; case x; unfold pred.\n";
+ for i = 0 to size do
+ fprintf fmt " intros x1 H1; unfold w%i_pred_c; \n" i;
+ fprintf fmt " generalize (spec_pred_c w%i_spec x1); case znz_pred_c; intros y1.\n" i;
+ fprintf fmt " rewrite spec_reduce_%i; auto.\n" i;
+ fprintf fmt " unfold interp_carry; unfold to_Z.\n";
+ fprintf fmt " case (spec_to_Z w%i_spec x1); intros HH1 HH2.\n" i;
+ fprintf fmt " case (spec_to_Z w%i_spec y1); intros HH3 HH4 HH5.\n" i;
+ fprintf fmt " assert (znz_to_Z w%i_op x1 - 1 < 0); auto with zarith.\n" i;
+ fprintf fmt " unfold to_Z in H1; auto with zarith.\n";
+ done;
+ fprintf fmt " intros n x1 H1; \n";
+ fprintf fmt " generalize (spec_pred_c (wn_spec n) x1); case znz_pred_c; intros y1.\n";
+ fprintf fmt " rewrite spec_reduce_n; auto.\n";
+ fprintf fmt " unfold interp_carry; unfold to_Z.\n";
+ fprintf fmt " case (spec_to_Z (wn_spec n) x1); intros HH1 HH2.\n";
+ fprintf fmt " case (spec_to_Z (wn_spec n) y1); intros HH3 HH4 HH5.\n";
+ fprintf fmt " assert (znz_to_Z (make_op n) x1 - 1 < 0); auto with zarith.\n";
+ fprintf fmt " unfold to_Z in H1; auto with zarith.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt " \n";
+
+ fprintf fmt " Let spec_pred0: forall x, [x] = 0 -> [pred x] = 0.\n";
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros x; case x; unfold pred.\n";
+ for i = 0 to size do
+ fprintf fmt " intros x1 H1; unfold w%i_pred_c; \n" i;
+ fprintf fmt " generalize (spec_pred_c w%i_spec x1); case znz_pred_c; intros y1.\n" i;
+ fprintf fmt " unfold interp_carry; unfold to_Z.\n";
+ fprintf fmt " unfold to_Z in H1; auto with zarith.\n";
+ fprintf fmt " case (spec_to_Z w%i_spec y1); intros HH3 HH4; auto with zarith.\n" i;
+ fprintf fmt " intros; exact (spec_0 w0_spec).\n";
+ done;
+ fprintf fmt " intros n x1 H1; \n";
+ fprintf fmt " generalize (spec_pred_c (wn_spec n) x1); case znz_pred_c; intros y1.\n";
+ fprintf fmt " unfold interp_carry; unfold to_Z.\n";
+ fprintf fmt " unfold to_Z in H1; auto with zarith.\n";
+ fprintf fmt " case (spec_to_Z (wn_spec n) y1); intros HH3 HH4; auto with zarith.\n";
+ fprintf fmt " intros; exact (spec_0 w0_spec).\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt " \n";
+
+
+ fprintf fmt " (***************************************************************)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (* Subtraction *)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (***************************************************************)\n\n";
+
for i = 0 to size do
fprintf fmt " Definition w%i_sub_c := w%i_op.(znz_sub_c).\n" i i
done;
@@ -334,52 +1235,108 @@ let print_Make () =
fprintf fmt " end.\n";
fprintf fmt "\n";
- fprintf fmt " Definition sub x y :=\n";
- fprintf fmt " match x, y with\n";
- fprintf fmt " | %s0 wx, %s0 wy => w0_sub wx wy \n" c c;
- for j = 1 to size do
- fprintf fmt " | %s0 wx, %s%i wy =>\n" c c j;
- fprintf fmt " if w0_eq0 wx then zero else w%i_sub " j;
- if j = 1 then fprintf fmt "(WW w_0 wx) wy\n"
- else fprintf fmt "(extend%i w0 (WW w_0 wx)) wy\n" (j-1)
- done;
- fprintf fmt " | %s0 wx, %sn n wy =>\n" c c;
- fprintf fmt " if w0_eq0 wx then zero\n";
- fprintf fmt " else subn n (extend n w%i (extend%i w0 (WW w_0 wx))) wy\n"
- size size;
- for i = 1 to size do
- fprintf fmt " | %s%i wx, %s0 wy =>" c i c;
- fprintf fmt "\n if w0_eq0 wy then x\n";
- fprintf fmt " else w%i_sub wx " i;
- if i = 1 then fprintf fmt "(WW w_0 wy)\n"
- else fprintf fmt "(extend%i w0 (WW w_0 wy))\n" (i-1);
- for j = 1 to size do
- fprintf fmt " | %s%i wx, %s%i wy => " c i c j;
- if i < j then fprintf fmt "w%i_sub (extend%i w%i wx) wy\n" j (j-i) (i-1)
- else if i = j then fprintf fmt "w%i_sub wx wy\n" j
- else fprintf fmt "w%i_sub wx (extend%i w%i wy)\n" i (i-j) (j-1)
- done;
- fprintf fmt
- " | %s%i wx, %sn n wy => subn n (extend n w%i (extend%i w%i wx)) wy\n"
- c i c size (size-i+1) (i-1)
- done;
- fprintf fmt " | %sn n wx, %s0 wy =>\n" c c;
- fprintf fmt " if w0_eq0 wy then x\n";
- fprintf fmt " else subn n wx (extend n w%i (extend%i w0 (WW w_0 wy)))\n"
- size size;
- for j = 1 to size do
- fprintf fmt
- " | %sn n wx, %s%i wy => subn n wx (extend n w%i (extend%i w%i wy))\n"
- c c j size (size-j+1) (j-1);
- done;
- fprintf fmt " | %sn n wx, %sn m wy =>\n" c c;
- fprintf fmt " let mn := Max.max n m in\n";
- fprintf fmt " let d := diff n m in\n";
- fprintf fmt " subn mn\n";
- fprintf fmt " (castm (diff_r n m) (extend_tr wx (snd d)))\n";
- fprintf fmt " (castm (diff_l n m) (extend_tr wy (fst d)))\n";
- fprintf fmt " end.\n";
- fprintf fmt "\n";
+ if gen_proof then
+ begin
+ for i = 0 to size do
+ fprintf fmt " Let spec_w%i_sub: forall x y, [%s%i y] <= [%s%i x] -> [w%i_sub x y] = [%s%i x] - [%s%i y].\n" i c i c i i c i c i;
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros n m; unfold w%i_sub, w%i_sub_c.\n" i i;
+ fprintf fmt " generalize (spec_sub_c w%i_spec n m); case znz_sub_c; \n" i;
+ if i == 0 then
+ fprintf fmt " intros x; auto.\n"
+ else
+ fprintf fmt " intros x; try rewrite spec_reduce_%i; auto.\n" i;
+ fprintf fmt " unfold interp_carry; unfold zero, w_0, to_Z.\n";
+ fprintf fmt " rewrite (spec_0 w0_spec).\n";
+ fprintf fmt " case (spec_to_Z w%i_spec x); intros; auto with zarith.\n" i;
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+ done;
+
+ fprintf fmt " Let spec_wn_sub: forall n x y, [%sn n y] <= [%sn n x] -> [subn n x y] = [%sn n x] - [%sn n y].\n" c c c c;
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros k n m; unfold subn.\n";
+ fprintf fmt " generalize (spec_sub_c (wn_spec k) n m); case znz_sub_c; \n";
+ fprintf fmt " intros x; auto.\n";
+ fprintf fmt " unfold interp_carry, to_Z.\n";
+ fprintf fmt " case (spec_to_Z (wn_spec k) x); intros; auto with zarith.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+ end;
+
+ fprintf fmt " Definition sub := Eval lazy beta delta [same_level] in\n";
+ fprintf fmt " (same_level t_ ";
+ for i = 0 to size do
+ fprintf fmt "w%i_sub " i;
+ done;
+ fprintf fmt "subn).\n";
+ fprintf fmt "\n";
+
+ fprintf fmt " Theorem spec_sub: forall x y, [y] <= [x] -> [sub x y] = [x] - [y].\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " unfold sub.\n";
+ fprintf fmt " generalize (spec_same_level t_ (fun x y res => y <= x -> [res] = x - y)).\n";
+ fprintf fmt " unfold same_level; intros HH; apply HH; clear HH.\n";
+ for i = 0 to size do
+ fprintf fmt " exact spec_w%i_sub.\n" i;
+ done;
+ fprintf fmt " exact spec_wn_sub.\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+
+ if gen_proof then
+ begin
+ for i = 0 to size do
+ fprintf fmt " Let spec_w%i_sub0: forall x y, [%s%i x] < [%s%i y] -> [w%i_sub x y] = 0.\n" i c i c i i;
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros n m; unfold w%i_sub, w%i_sub_c.\n" i i;
+ fprintf fmt " generalize (spec_sub_c w%i_spec n m); case znz_sub_c; \n" i;
+ fprintf fmt " intros x; unfold interp_carry.\n";
+ fprintf fmt " unfold to_Z; case (spec_to_Z w%i_spec x); intros; auto with zarith.\n" i;
+ fprintf fmt " intros; unfold to_Z, zero, w_0; rewrite (spec_0 w0_spec); auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+ done;
+
+ fprintf fmt " Let spec_wn_sub0: forall n x y, [%sn n x] < [%sn n y] -> [subn n x y] = 0.\n" c c;
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros k n m; unfold subn.\n";
+ fprintf fmt " generalize (spec_sub_c (wn_spec k) n m); case znz_sub_c; \n";
+ fprintf fmt " intros x; unfold interp_carry.\n";
+ fprintf fmt " unfold to_Z; case (spec_to_Z (wn_spec k) x); intros; auto with zarith.\n";
+ fprintf fmt " intros; unfold to_Z, w_0; rewrite (spec_0 (w0_spec)); auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+ end;
+
+ fprintf fmt " Theorem spec_sub0: forall x y, [x] < [y] -> [sub x y] = 0.\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " unfold sub.\n";
+ fprintf fmt " generalize (spec_same_level t_ (fun x y res => x < y -> [res] = 0)).\n";
+ fprintf fmt " unfold same_level; intros HH; apply HH; clear HH.\n";
+ for i = 0 to size do
+ fprintf fmt " exact spec_w%i_sub0.\n" i;
+ done;
+ fprintf fmt " exact spec_wn_sub0.\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+
+
+ fprintf fmt " (***************************************************************)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (* Comparison *)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (***************************************************************)\n\n";
for i = 0 to size do
fprintf fmt " Definition compare_%i := w%i_op.(znz_compare).\n" i i;
@@ -391,37 +1348,109 @@ let print_Make () =
done;
fprintf fmt "\n";
- (* Comparison *)
- fprintf fmt " Definition compare x y :=\n";
- fprintf fmt " match x, y with\n";
- for i = 0 to size do
- for j = 0 to size do
- fprintf fmt " | %s%i wx, %s%i wy => " c i c j;
- if i < j then fprintf fmt "opp_compare (comparen_%i %i wy wx)\n" i (j-i)
- else if i = j then fprintf fmt "compare_%i wx wy\n" i
- else fprintf fmt "comparen_%i %i wx wy\n" j (i-j)
- done;
- let s0 = if i = 0 then "w_0" else "W0" in
- fprintf fmt " | %s%i wx, %sn n wy =>\n" c i c;
- fprintf fmt " opp_compare (compare_mn_1 w%i w%i %s " size i s0;
- fprintf fmt "compare_%i (compare_%i W0) (comparen_%i %i) (S n) wy wx)\n"
- i size i (size - i)
- done;
- for j = 0 to size do
- let s0 = if j = 0 then "w_0" else "W0" in
- fprintf fmt " | %sn n wx, %s%i wy =>\n" c c j;
- fprintf fmt " compare_mn_1 w%i w%i %s " size j s0;
- fprintf fmt "compare_%i (compare_%i W0) (comparen_%i %i) (S n) wx wy\n"
- j size j (size - j)
- done;
- fprintf fmt " | %sn n wx, %sn m wy =>\n" c c;
+ fprintf fmt " Definition comparenm n m wx wy :=\n";
fprintf fmt " let mn := Max.max n m in\n";
fprintf fmt " let d := diff n m in\n";
fprintf fmt " let op := make_op mn in\n";
fprintf fmt " op.(znz_compare)\n";
fprintf fmt " (castm (diff_r n m) (extend_tr wx (snd d)))\n";
- fprintf fmt " (castm (diff_l n m) (extend_tr wy (fst d)))\n";
- fprintf fmt " end.\n";
+ fprintf fmt " (castm (diff_l n m) (extend_tr wy (fst d))).\n";
+ fprintf fmt "\n";
+
+ fprintf fmt " Definition compare := Eval lazy beta delta [iter] in \n";
+ fprintf fmt " (iter _ \n";
+ for i = 0 to size do
+ fprintf fmt " compare_%i\n" i;
+ fprintf fmt " (fun n x y => opp_compare (comparen_%i (S n) y x))\n" i;
+ fprintf fmt " (fun n => comparen_%i (S n))\n" i;
+ done;
+ fprintf fmt " comparenm).\n";
+ fprintf fmt "\n";
+
+ if gen_proof then
+ begin
+ for i = 0 to size do
+ fprintf fmt " Let spec_compare_%i: forall x y,\n" i;
+ fprintf fmt " match compare_%i x y with \n" i;
+ fprintf fmt " Eq => [%s%i x] = [%s%i y]\n" c i c i;
+ fprintf fmt " | Lt => [%s%i x] < [%s%i y]\n" c i c i;
+ fprintf fmt " | Gt => [%s%i x] > [%s%i y]\n" c i c i;
+ fprintf fmt " end.\n";
+ fprintf fmt " Proof.\n";
+ fprintf fmt " unfold compare_%i, to_Z; exact (spec_compare w%i_spec).\n" i i;
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+
+ fprintf fmt " Let spec_comparen_%i:\n" i;
+ fprintf fmt " forall (n : nat) (x : word w%i n) (y : w%i),\n" i i;
+ fprintf fmt " match comparen_%i n x y with\n" i;
+ fprintf fmt " | Eq => eval%in n x = [%s%i y]\n" i c i;
+ fprintf fmt " | Lt => eval%in n x < [%s%i y]\n" i c i;
+ fprintf fmt " | Gt => eval%in n x > [%s%i y]\n" i c i;
+ fprintf fmt " end.\n";
+ fprintf fmt " intros n x y.\n";
+ fprintf fmt " unfold comparen_%i, to_Z; rewrite spec_gen_eval%in.\n" i i;
+ fprintf fmt " apply spec_compare_mn_1.\n";
+ fprintf fmt " exact (spec_0 w%i_spec).\n" i;
+ if i == 0 then
+ fprintf fmt " intros x1; exact (spec_compare w%i_spec w_0 x1).\n" i
+ else
+ fprintf fmt " intros x1; exact (spec_compare w%i_spec W0 x1).\n" i;
+ fprintf fmt " exact (spec_to_Z w%i_spec).\n" i;
+ fprintf fmt " exact (spec_compare w%i_spec).\n" i;
+ fprintf fmt " exact (spec_compare w%i_spec).\n" i;
+ fprintf fmt " exact (spec_to_Z w%i_spec).\n" i;
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+
+ done;
+
+ fprintf fmt " Let spec_opp_compare: forall c (u v: Z),\n";
+ fprintf fmt " match c with Eq => u = v | Lt => u < v | Gt => u > v end ->\n";
+ fprintf fmt " match opp_compare c with Eq => v = u | Lt => v < u | Gt => v > u end.\n";
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros c u v; case c; unfold opp_compare; auto with zarith.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+ end;
+
+ fprintf fmt " Theorem spec_compare: forall x y,\n";
+ fprintf fmt " match compare x y with \n";
+ fprintf fmt " Eq => [x] = [y]\n";
+ fprintf fmt " | Lt => [x] < [y]\n";
+ fprintf fmt " | Gt => [x] > [y]\n";
+ fprintf fmt " end.\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " refine (spec_iter _ (fun x y res => \n";
+ fprintf fmt " match res with \n";
+ fprintf fmt " Eq => x = y\n";
+ fprintf fmt " | Lt => x < y\n";
+ fprintf fmt " | Gt => x > y\n";
+ fprintf fmt " end)\n";
+ for i = 0 to size do
+ fprintf fmt " compare_%i\n" i;
+ fprintf fmt " (fun n x y => opp_compare (comparen_%i (S n) y x))\n" i;
+ fprintf fmt " (fun n => comparen_%i (S n)) _ _ _\n" i;
+ done;
+ fprintf fmt " comparenm _).\n";
+
+ for i = 0 to size - 1 do
+ fprintf fmt " exact spec_compare_%i.\n" i;
+ fprintf fmt " intros n x y H;apply spec_opp_compare; apply spec_comparen_%i.\n" i;
+ fprintf fmt " intros n x y H; exact (spec_comparen_%i (S n) x y).\n" i;
+ done;
+ fprintf fmt " exact spec_compare_%i.\n" size;
+ fprintf fmt " intros n x y;apply spec_opp_compare; apply spec_comparen_%i.\n" size;
+ fprintf fmt " intros n; exact (spec_comparen_%i (S n)).\n" size;
+ fprintf fmt " intros n m x y; unfold comparenm.\n";
+ fprintf fmt " rewrite <- (spec_cast_l n m x); rewrite <- (spec_cast_r n m y).\n";
+ fprintf fmt " unfold to_Z; apply (spec_compare (wn_spec (Max.max n m))).\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
fprintf fmt "\n";
fprintf fmt " Definition eq_bool x y :=\n";
@@ -430,9 +1459,28 @@ let print_Make () =
fprintf fmt " | _ => false\n";
fprintf fmt " end.\n";
fprintf fmt "\n";
-
-
- (* Multiplication *)
+
+
+ fprintf fmt " Theorem spec_eq_bool: forall x y,\n";
+ fprintf fmt " if eq_bool x y then [x] = [y] else [x] <> [y].\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros x y; unfold eq_bool.\n";
+ fprintf fmt " generalize (spec_compare x y); case compare; auto with zarith.\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+
+
+
+ fprintf fmt " (***************************************************************)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (* Multiplication *)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (***************************************************************)\n\n";
for i = 0 to size do
fprintf fmt " Definition w%i_mul_c := w%i_op.(znz_mul_c).\n" i i
done;
@@ -448,6 +1496,11 @@ let print_Make () =
fprintf fmt "\n";
for i = 0 to size do
+ fprintf fmt " Definition w%i_0W := w%i_op.(znz_0W).\n" i i
+ done;
+ fprintf fmt "\n";
+
+ for i = 0 to size do
let s0 = if i = 0 then "w_0" else "W0" in
fprintf fmt " Definition w%i_mul_add_n1 :=\n" i;
fprintf fmt
@@ -456,103 +1509,234 @@ let print_Make () =
done;
fprintf fmt "\n";
- fprintf fmt " Definition mul x y :=\n";
- fprintf fmt " match x, y with\n";
- fprintf fmt " | %s0 wx, %s0 wy =>\n" c c;
- fprintf fmt " reduce_1 (w0_mul_c wx wy)\n";
- for j = 1 to size do
- fprintf fmt " | %s0 wx, %s%i wy =>\n" c c j;
- fprintf fmt " if w0_eq0 wx then zero\n";
- fprintf fmt " else\n";
- fprintf fmt " let (w,r) := w0_mul_add_n1 %i wy wx w_0 in\n" j;
- fprintf fmt " if w0_eq0 w then %s%i r\n" c j;
- if j = 1 then
- fprintf fmt " else %s2 (WW (WW w_0 w) r)\n" c
- else if j = size then
- fprintf fmt " else %sn 0 (WW (extend%i w0 (WW w_0 w)) r)\n"
- c (size-1)
- else
- fprintf fmt " else %s%i (WW (extend%i w0 (WW w_0 w)) r)\n"
- c (j+1) (j-1)
- done;
+ begin
+ for i = 0 to size - 1 do
+ fprintf fmt " Let to_Z%i n :=\n" i;
+ fprintf fmt " match n return word w%i (S n) -> t_ with\n" i;
+ for j = 0 to size - i do
+ if (i + j) == size then
+ begin
+ fprintf fmt " | %i%s => fun x => %sn 0 x\n" j "%nat" c;
+ fprintf fmt " | %i%s => fun x => %sn 1 x\n" (j + 1) "%nat" c
+ end
+ else
+ fprintf fmt " | %i%s => fun x => %s%i x\n" j "%nat" c (i + j + 1)
+ done;
+ fprintf fmt " | _ => fun _ => N0 w_0\n";
+ fprintf fmt " end.\n";
+ fprintf fmt "\n";
+ done;
- fprintf fmt " | %s0 wx, %sn n wy =>\n" c c;
- fprintf fmt " if w0_eq0 wx then zero\n";
- fprintf fmt " else\n";
- fprintf fmt " let (w,r) := w%i_mul_add_n1 (S n) wy " size;
- fprintf fmt "(extend%i w0 (WW w_0 wx)) W0 in\n" (size - 1);
- fprintf fmt " if w%i_eq0 w then %sn n r\n" size c;
- fprintf fmt " else %sn (S n) (WW (extend n w%i (WW W0 w)) r)\n" c size;
-
- for i = 1 to size do
- fprintf fmt " | %s%i wx, %s0 wy =>\n" c i c;
- fprintf fmt " if w0_eq0 wy then zero\n";
- fprintf fmt " else\n";
- fprintf fmt " let (w,r) := w0_mul_add_n1 %i wx wy w_0 in\n" i;
- fprintf fmt " if w0_eq0 w then %s%i r\n" c i;
- if i = 1 then
- fprintf fmt " else %s2 (WW (WW w_0 w) r)\n" c
- else if i = size then
- fprintf fmt " else %sn 0 (WW (extend%i w0 (WW w_0 w)) r)\n"
- c (size-1)
- else
- fprintf fmt " else %s%i (WW (extend%i w0 (WW w_0 w)) r)\n"
- c (i+1) (i-1);
- for j = 1 to size do
- fprintf fmt " | %s%i wx, %s%i wy =>\n" c i c j;
- if i = j then begin
- if i = size then fprintf fmt " %sn 0 (w%i_mul_c wx wy)\n" c i
- else fprintf fmt " %s%i (w%i_mul_c wx wy)\n" c (i+1) i
- end else begin
- let min,max, wmin, wmax =
- if i < j then i, j, "wx", "wy" else j, i, "wy", "wx" in
- fprintf fmt " let (w,r) := w%i_mul_add_n1 %i %s %s W0 in\n"
- min (max-min) wmax wmin;
- fprintf fmt " if w%i_eq0 w then %s%i r\n" min c max;
- fprintf fmt " else ";
- if max = size then fprintf fmt "%sn 0 " c
- else fprintf fmt "%s%i " c (max+1);
- fprintf fmt "(WW (extend%i w%i w) r)\n" (max - min) (min-1);
- end
+
+ if gen_proof then
+ for i = 0 to size - 1 do
+ fprintf fmt "Theorem to_Z%i_spec:\n" i;
+ fprintf fmt " forall n x, Z_of_nat n <= %i -> [to_Z%i n x] = znz_to_Z (nmake_op _ w%i_op (S n)) x.\n" (size + 1 - i) i i;
+ for j = 1 to size + 2 - i do
+ fprintf fmt " intros n; case n; clear n.\n";
+ fprintf fmt " unfold to_Z%i.\n" i;
+ fprintf fmt " intros x H; rewrite spec_eval%in%i; auto.\n" i j;
done;
- fprintf fmt " | %s%i wx, %sn n wy =>\n" c i c;
- fprintf fmt " let (w,r) := w%i_mul_add_n1 (S n) wy " size;
- if i = size then fprintf fmt "wx W0 in\n"
- else
- fprintf fmt "(extend%i w%i wx) W0 in\n" (size - i) (i-1);
- fprintf fmt " if w%i_eq0 w then %sn n r\n" size c;
- fprintf fmt " else %sn (S n) (WW (extend n w%i (WW W0 w)) r)\n" c size
-
- done;
- fprintf fmt " | %sn n wx, %s0 wy =>\n" c c;
- fprintf fmt " if w0_eq0 wy then zero\n";
- fprintf fmt " else\n";
- fprintf fmt " let (w,r) := w%i_mul_add_n1 (S n) wx " size;
- fprintf fmt "(extend%i w0 (WW w_0 wy)) W0 in\n" (size - 1);
- fprintf fmt " if w%i_eq0 w then %sn n r\n" size c;
- fprintf fmt " else %sn (S n) (WW (extend n w%i (WW W0 w)) r)\n" c size;
-
- for j = 1 to size do
- fprintf fmt " | %sn n wx, %s%i wy =>\n" c c j;
- fprintf fmt " let (w,r) := w%i_mul_add_n1 (S n) wx " size;
- if j = size then fprintf fmt "wy W0 in\n"
+ fprintf fmt " intros n x.\n";
+ fprintf fmt " repeat rewrite inj_S; unfold Zsucc; auto with zarith.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+ done;
+ end;
+
+ for i = 0 to size do
+ fprintf fmt " Definition w%i_mul n x y :=\n" i;
+ if i == 0 then
+ fprintf fmt " let (w,r) := w%i_mul_add_n1 (S n) x y w_0 in\n" i
else
- fprintf fmt "(extend%i w%i wy) W0 in\n" (size - j) (j-1);
- fprintf fmt " if w%i_eq0 w then %sn n r\n" size c;
- fprintf fmt " else %sn (S n) (WW (extend n w%i (WW W0 w)) r)\n" c size
+ fprintf fmt " let (w,r) := w%i_mul_add_n1 (S n) x y W0 in\n" i;
+ if i == size then
+ begin
+ fprintf fmt " if w%i_eq0 w then %sn n r\n" i c;
+ fprintf fmt " else %sn (S n) (WW (extend%i n w) r).\n" c i;
+ end
+ else
+ begin
+ fprintf fmt " if w%i_eq0 w then to_Z%i n r\n" i i;
+ fprintf fmt " else to_Z%i (S n) (WW (extend%i n w) r).\n" i i;
+ end;
+ fprintf fmt "\n";
done;
- fprintf fmt " | %sn n wx, %sn m wy =>\n" c c;
+ fprintf fmt " Definition mulnm n m x y :=\n";
fprintf fmt " let mn := Max.max n m in\n";
fprintf fmt " let d := diff n m in\n";
fprintf fmt " let op := make_op mn in\n";
fprintf fmt " reduce_n (S mn) (op.(znz_mul_c)\n";
- fprintf fmt " (castm (diff_r n m) (extend_tr wx (snd d)))\n";
- fprintf fmt " (castm (diff_l n m) (extend_tr wy (fst d))))\n";
- fprintf fmt " end.\n";
+ fprintf fmt " (castm (diff_r n m) (extend_tr x (snd d)))\n";
+ fprintf fmt " (castm (diff_l n m) (extend_tr y (fst d)))).\n";
fprintf fmt "\n";
-
- (* Square *)
+
+ fprintf fmt " Definition mul := Eval lazy beta delta [iter0] in \n";
+ fprintf fmt " (iter0 t_ \n";
+ for i = 0 to size do
+ fprintf fmt " (fun x y => reduce_%i (w%i_mul_c x y)) \n" (i + 1) i;
+ fprintf fmt " (fun n x y => w%i_mul n y x)\n" i;
+ fprintf fmt " w%i_mul\n" i;
+ done;
+ fprintf fmt " mulnm\n";
+ fprintf fmt " (fun _ => N0 w_0)\n";
+ fprintf fmt " (fun _ => N0 w_0)\n";
+ fprintf fmt " ).\n";
+ fprintf fmt "\n";
+ if gen_proof then
+ begin
+ for i = 0 to size do
+ fprintf fmt " Let spec_w%i_mul_add: forall x y z,\n" i;
+ fprintf fmt " let (q,r) := w%i_mul_add x y z in\n" i;
+ fprintf fmt " znz_to_Z w%i_op q * (base (znz_digits w%i_op)) + znz_to_Z w%i_op r =\n" i i i;
+ fprintf fmt " znz_to_Z w%i_op x * znz_to_Z w%i_op y + znz_to_Z w%i_op z :=\n" i i i ;
+ fprintf fmt " (spec_mul_add w%i_spec).\n" i;
+ fprintf fmt "\n";
+ done;
+
+ for i = 0 to size do
+
+
+ fprintf fmt " Theorem spec_w%i_mul_add_n1: forall n x y z,\n" i;
+ fprintf fmt " let (q,r) := w%i_mul_add_n1 n x y z in\n" i;
+ fprintf fmt " znz_to_Z w%i_op q * (base (znz_digits (nmake_op _ w%i_op n))) +\n" i i;
+ fprintf fmt " znz_to_Z (nmake_op _ w%i_op n) r =\n" i;
+ fprintf fmt " znz_to_Z (nmake_op _ w%i_op n) x * znz_to_Z w%i_op y +\n" i i;
+ fprintf fmt " znz_to_Z w%i_op z.\n" i;
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros n x y z; unfold w%i_mul_add_n1.\n" i;
+ fprintf fmt " rewrite nmake_gen.\n";
+ fprintf fmt " rewrite digits_gend.\n";
+ fprintf fmt " change (base (GenBase.gen_digits (znz_digits w%i_op) n)) with\n" i;
+ fprintf fmt " (GenBase.gen_wB (znz_digits w%i_op) n).\n" i;
+ fprintf fmt " apply spec_gen_mul_add_n1; auto.\n";
+ if i == 0 then fprintf fmt " exact (spec_0 w%i_spec).\n" i;
+ fprintf fmt " exact (spec_WW w%i_spec).\n" i;
+ fprintf fmt " exact (spec_0W w%i_spec).\n" i;
+ fprintf fmt " exact (spec_mul_add w%i_spec).\n" i;
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+ done;
+
+ fprintf fmt " Lemma nmake_op_WW: forall ww ww1 n x y,\n";
+ fprintf fmt " znz_to_Z (nmake_op ww ww1 (S n)) (WW x y) =\n";
+ fprintf fmt " znz_to_Z (nmake_op ww ww1 n) x * base (znz_digits (nmake_op ww ww1 n)) +\n";
+ fprintf fmt " znz_to_Z (nmake_op ww ww1 n) y.\n";
+ fprintf fmt " auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+
+ for i = 0 to size do
+ fprintf fmt " Lemma extend%in_spec: forall n x1,\n" i;
+ fprintf fmt " znz_to_Z (nmake_op _ w%i_op (S n)) (extend%i n x1) = \n" i i;
+ fprintf fmt " znz_to_Z w%i_op x1.\n" i;
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros n1 x2; rewrite nmake_gen.\n";
+ fprintf fmt " unfold extend%i.\n" i;
+ fprintf fmt " rewrite GenBase.spec_extend; auto.\n";
+ if (i == 0) then
+ fprintf fmt " intros l; simpl; unfold w_0; rewrite (spec_0 w0_spec); ring.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+
+ done;
+
+ fprintf fmt " Lemma spec_muln:\n";
+ fprintf fmt " forall n (x: word _ (S n)) y,\n";
+ fprintf fmt " [%sn (S n) (znz_mul_c (make_op n) x y)] = [%sn n x] * [%sn n y].\n" c c c;
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros n x y; unfold to_Z.\n";
+ fprintf fmt " rewrite <- (spec_mul_c (wn_spec n)).\n";
+ fprintf fmt " rewrite make_op_S.\n";
+ fprintf fmt " case znz_mul_c; auto.\n";
+ fprintf fmt " Qed.\n";
+ end;
+
+ fprintf fmt " Theorem spec_mul: forall x y, [mul x y] = [x] * [y].\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ for i = 0 to size do
+ fprintf fmt " assert(F%i: \n" i;
+ fprintf fmt " forall n x y,\n";
+ if i <> size then
+ fprintf fmt " Z_of_nat n <= %i -> " (size - i);
+ fprintf fmt " [w%i_mul n x y] = eval%in (S n) x * [%s%i y]).\n" i i c i;
+ if i == size then
+ fprintf fmt " intros n x y; unfold w%i_mul.\n" i
+ else
+ fprintf fmt " intros n x y H; unfold w%i_mul.\n" i;
+ if i == 0 then
+ fprintf fmt " generalize (spec_w%i_mul_add_n1 (S n) x y w_0).\n" i
+ else
+ fprintf fmt " generalize (spec_w%i_mul_add_n1 (S n) x y W0).\n" i;
+ fprintf fmt " case w%i_mul_add_n1; intros x1 y1.\n" i;
+ fprintf fmt " change (znz_to_Z (nmake_op _ w%i_op (S n)) x) with (eval%in (S n) x).\n" i i;
+ fprintf fmt " change (znz_to_Z w%i_op y) with ([%s%i y]).\n" i c i;
+ if i == 0 then
+ fprintf fmt " unfold w_0; rewrite (spec_0 w0_spec); rewrite Zplus_0_r.\n"
+ else
+ fprintf fmt " change (znz_to_Z w%i_op W0) with 0; rewrite Zplus_0_r.\n" i;
+ fprintf fmt " intros H1; rewrite <- H1; clear H1.\n";
+ fprintf fmt " generalize (spec_w%i_eq0 x1); case w%i_eq0; intros HH.\n" i i;
+ fprintf fmt " unfold to_Z in HH; rewrite HH.\n";
+ if i == size then
+ begin
+ fprintf fmt " rewrite spec_eval%in; unfold eval%in, nmake_op%i; auto.\n" i i i;
+ fprintf fmt " rewrite spec_eval%in; unfold eval%in, nmake_op%i.\n" i i i
+ end
+ else
+ begin
+ fprintf fmt " rewrite to_Z%i_spec; auto with zarith.\n" i;
+ fprintf fmt " rewrite to_Z%i_spec; try (rewrite inj_S; auto with zarith).\n" i
+ end;
+ fprintf fmt " rewrite nmake_op_WW; rewrite extend%in_spec; auto.\n" i;
+ done;
+ fprintf fmt " refine (spec_iter0 t_ (fun x y res => [res] = x * y)\n";
+ for i = 0 to size do
+ fprintf fmt " (fun x y => reduce_%i (w%i_mul_c x y)) \n" (i + 1) i;
+ fprintf fmt " (fun n x y => w%i_mul n y x)\n" i;
+ fprintf fmt " w%i_mul _ _ _\n" i;
+ done;
+ fprintf fmt " mulnm _\n";
+ fprintf fmt " (fun _ => N0 w_0) _\n";
+ fprintf fmt " (fun _ => N0 w_0) _\n";
+ fprintf fmt " ).\n";
+ for i = 0 to size do
+ fprintf fmt " intros x y; rewrite spec_reduce_%i.\n" (i + 1);
+ fprintf fmt " unfold w%i_mul_c, to_Z.\n" i;
+ fprintf fmt " generalize (spec_mul_c w%i_spec x y).\n" i;
+ fprintf fmt " intros HH; rewrite <- HH; clear HH; auto.\n";
+ if i == size then
+ begin
+ fprintf fmt " intros n x y; rewrite F%i; auto with zarith.\n" i;
+ fprintf fmt " intros n x y; rewrite F%i; auto with zarith. \n" i;
+ end
+ else
+ begin
+ fprintf fmt " intros n x y H; rewrite F%i; auto with zarith.\n" i;
+ fprintf fmt " intros n x y H; rewrite F%i; auto with zarith. \n" i;
+ end;
+ done;
+ fprintf fmt " intros n m x y; unfold mulnm.\n";
+ fprintf fmt " rewrite spec_reduce_n.\n";
+ fprintf fmt " rewrite <- (spec_cast_l n m x).\n";
+ fprintf fmt " rewrite <- (spec_cast_r n m y).\n";
+ fprintf fmt " rewrite spec_muln; rewrite spec_cast_l; rewrite spec_cast_r; auto.\n";
+ fprintf fmt " intros x; unfold to_Z, w_0; rewrite (spec_0 w0_spec); ring.\n";
+ fprintf fmt " intros x; unfold to_Z, w_0; rewrite (spec_0 w0_spec); ring.\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+
+ fprintf fmt " (***************************************************************)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (* Square *)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (***************************************************************)\n\n";
for i = 0 to size do
fprintf fmt " Definition w%i_square_c := w%i_op.(znz_square_c).\n" i i
done;
@@ -571,6 +1755,32 @@ let print_Make () =
fprintf fmt " end.\n";
fprintf fmt "\n";
+ fprintf fmt " Theorem spec_square: forall x, [square x] = [x] * [x].\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros x; case x; unfold square; clear x.\n";
+ fprintf fmt " intros x; rewrite spec_reduce_1; unfold to_Z.\n";
+ fprintf fmt " exact (spec_square_c w%i_spec x).\n" 0;
+ for i = 1 to size do
+ fprintf fmt " intros x; unfold to_Z.\n";
+ fprintf fmt " exact (spec_square_c w%i_spec x).\n" i;
+ done;
+ fprintf fmt " intros n x; unfold to_Z.\n";
+ fprintf fmt " rewrite make_op_S.\n";
+ fprintf fmt " exact (spec_square_c (wn_spec n) x).\n";
+ fprintf fmt "Qed.\n";
+ end
+ else
+ fprintf fmt "Admitted.\n";
+ fprintf fmt "\n";
+
+
+ fprintf fmt " (***************************************************************)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (* Power *)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (***************************************************************)\n\n";
fprintf fmt " Fixpoint power_pos (x:%s) (p:positive) {struct p} : %s :=\n"
t t;
fprintf fmt " match p with\n";
@@ -579,13 +1789,39 @@ let print_Make () =
fprintf fmt " | xI p => mul (square (power_pos x p)) x\n";
fprintf fmt " end.\n";
fprintf fmt "\n";
-
- (* Square root *)
+
+ fprintf fmt " Theorem spec_power_pos: forall x n, [power_pos x n] = [x] ^ Zpos n.\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros x n; generalize x; elim n; clear n x; simpl power_pos.\n";
+ fprintf fmt " intros; rewrite spec_mul; rewrite spec_square; rewrite H.\n";
+ fprintf fmt " rewrite Zpos_xI; rewrite Zpower_exp; auto with zarith.\n";
+ fprintf fmt " rewrite (Zmult_comm 2); rewrite ZPowerAux.Zpower_mult; auto with zarith.\n";
+ fprintf fmt " rewrite ZAux.Zpower_2; rewrite ZAux.Zpower_exp_1; auto.\n";
+ fprintf fmt " intros; rewrite spec_square; rewrite H.\n";
+ fprintf fmt " rewrite Zpos_xO; auto with zarith.\n";
+ fprintf fmt " rewrite (Zmult_comm 2); rewrite ZPowerAux.Zpower_mult; auto with zarith.\n";
+ fprintf fmt " rewrite ZAux.Zpower_2; auto.\n";
+ fprintf fmt " intros; rewrite ZAux.Zpower_exp_1; auto.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+
+ fprintf fmt " (***************************************************************)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (* Square root *)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (***************************************************************)\n\n";
+
for i = 0 to size do
fprintf fmt " Definition w%i_sqrt := w%i_op.(znz_sqrt).\n" i i
done;
fprintf fmt "\n";
-
+
fprintf fmt " Definition sqrt x :=\n";
fprintf fmt " match x with\n";
for i = 0 to size do
@@ -597,6 +1833,30 @@ let print_Make () =
fprintf fmt " end.\n";
fprintf fmt "\n";
+
+
+ fprintf fmt " Theorem spec_sqrt: forall x, [sqrt x] ^ 2 <= [x] < ([sqrt x] + 1) ^ 2.\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros x; unfold sqrt; case x; clear x.\n";
+ for i = 0 to size do
+ fprintf fmt " intros x; rewrite spec_reduce_%i; exact (spec_sqrt w%i_spec x).\n" i i;
+ done;
+ fprintf fmt " intros n x; rewrite spec_reduce_n; exact (spec_sqrt (wn_spec n) x).\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt "Admitted.\n";
+ fprintf fmt "\n";
+
+
+ fprintf fmt " (***************************************************************)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (* Division *)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (***************************************************************)\n\n";
+
(* Division *)
for i = 0 to size do
@@ -604,62 +1864,168 @@ let print_Make () =
done;
fprintf fmt "\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Let spec_divn1 ww (ww_op: znz_op ww) (ww_spec: znz_spec ww_op) := \n";
+ fprintf fmt " (spec_gen_divn1 \n";
+ fprintf fmt " ww_op.(znz_zdigits) ww_op.(znz_0)\n";
+ fprintf fmt " ww_op.(znz_WW) ww_op.(znz_head0)\n";
+ fprintf fmt " ww_op.(znz_add_mul_div) ww_op.(znz_div21)\n";
+ fprintf fmt " ww_op.(znz_compare) ww_op.(znz_sub) (znz_to_Z ww_op)\n";
+ fprintf fmt " (spec_to_Z ww_spec) \n";
+ fprintf fmt " (spec_zdigits ww_spec)\n";
+ fprintf fmt " (spec_0 ww_spec) (spec_WW ww_spec) (spec_head0 ww_spec)\n";
+ fprintf fmt " (spec_add_mul_div ww_spec) (spec_div21 ww_spec) \n";
+ fprintf fmt " (ZnZ.spec_compare ww_spec) (ZnZ.spec_sub ww_spec)).\n";
+ fprintf fmt " \n";
+ end;
+
for i = 0 to size do
- fprintf fmt " Definition w%i_divn1 :=\n" i;
+ fprintf fmt " Definition w%i_divn1 n x y :=\n" i;
+ fprintf fmt " let (u, v) :=\n";
fprintf fmt " gen_divn1 w%i_op.(znz_zdigits) w%i_op.(znz_0)\n" i i;
fprintf fmt " w%i_op.(znz_WW) w%i_op.(znz_head0)\n" i i;
fprintf fmt " w%i_op.(znz_add_mul_div) w%i_op.(znz_div21)\n" i i;
- fprintf fmt " w%i_op.(znz_compare) w%i_op.(znz_sub).\n" i i;
+ fprintf fmt " w%i_op.(znz_compare) w%i_op.(znz_sub) (S n) x y in\n" i i;
+ if i == size then
+ fprintf fmt " (%sn _ u, %s%i v).\n" c c i
+ else
+ fprintf fmt " (to_Z%i _ u, %s%i v).\n" i c i;
done;
fprintf fmt "\n";
- fprintf fmt " Definition div_gt x y :=\n";
- fprintf fmt " match x, y with\n";
+
+ if gen_proof then
+ begin
for i = 0 to size do
- for j = 0 to size do
- fprintf fmt " | %s%i wx, %s%i wy =>" c i c j;
- if i = j then
- fprintf fmt
- " let (q, r):= w%i_div_gt wx wy in (reduce_%i q, reduce_%i r)\n"
- i i i
- else if i > j then
- fprintf fmt
- " let (q, r):= w%i_divn1 %i wx wy in (reduce_%i q, reduce_%i r)\n"
- j (i-j) i j
- else begin (* i < j *)
- fprintf fmt
- "\n let wx':= GenBase.extend w%i_0W %i wx in\n"
- i (j-i-1);
- fprintf fmt " let (q, r):= w%i_div_gt wx' wy in\n" j;
- fprintf fmt " (reduce_%i q, reduce_%i r)\n" j j;
- end
- done;
- fprintf fmt " | %s%i wx, %sn n wy =>\n" c i c;
- fprintf fmt
- " let wx':= extend n w%i (GenBase.extend w%i_0W %i wx) in\n"
- size i (size-i);
- fprintf fmt " let (q, r):= (make_op n).(znz_div_gt) wx' wy in\n";
- fprintf fmt " (reduce_n n q, reduce_n n r)\n";
- done;
- for j = 0 to size do
- fprintf fmt " | %sn n wx, %s%i wy =>\n" c c j;
- if j < size then
- fprintf fmt " let wy':= GenBase.extend w%i_0W %i wy in\n"
- j (size-j-1)
- else
- fprintf fmt " let wy':= wy in\n";
- fprintf fmt " let (q, r):= w%i_divn1 (S n) wx wy' in\n" size;
- fprintf fmt " (reduce_n n q, reduce_%i r)\n" size
+ fprintf fmt " Lemma spec_get_end%i: forall n x y,\n" i;
+ fprintf fmt " eval%in n x <= [%s%i y] -> \n" i c i;
+ fprintf fmt " [%s%i (GenBase.get_low %s n x)] = eval%in n x.\n" c i (pz i) i;
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros n x y H.\n";
+ fprintf fmt " rewrite spec_gen_eval%in; unfold to_Z.\n" i;
+ fprintf fmt " apply GenBase.spec_get_low.\n";
+ fprintf fmt " exact (spec_0 w%i_spec).\n" i;
+ fprintf fmt " exact (spec_to_Z w%i_spec).\n" i;
+ fprintf fmt " apply Zle_lt_trans with [%s%i y]; auto.\n" c i;
+ fprintf fmt " rewrite <- spec_gen_eval%in; auto.\n" i;
+ fprintf fmt " unfold to_Z; case (spec_to_Z w%i_spec y); auto.\n" i;
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+ done ;
+ end;
+
+ for i = 0 to size do
+ fprintf fmt " Let div_gt%i x y := let (u,v) := (w%i_div_gt x y) in (reduce_%i u, reduce_%i v).\n" i i i i;
done;
- fprintf fmt " | %sn n wx, %sn m wy =>\n" c c;
+ fprintf fmt "\n";
+
+
+ fprintf fmt " Let div_gtnm n m wx wy :=\n";
fprintf fmt " let mn := Max.max n m in\n";
fprintf fmt " let d := diff n m in\n";
fprintf fmt " let op := make_op mn in\n";
- fprintf fmt " let (q, r):= op.(znz_div)\n";
+ fprintf fmt " let (q, r):= op.(znz_div_gt)\n";
fprintf fmt " (castm (diff_r n m) (extend_tr wx (snd d)))\n";
fprintf fmt " (castm (diff_l n m) (extend_tr wy (fst d))) in\n";
- fprintf fmt " (reduce_n mn q, reduce_n mn r)\n";
- fprintf fmt " end.\n";
+ fprintf fmt " (reduce_n mn q, reduce_n mn r).\n";
+ fprintf fmt "\n";
+
+ fprintf fmt " Definition div_gt := Eval lazy beta delta [iter] in\n";
+ fprintf fmt " (iter _ \n";
+ for i = 0 to size do
+ fprintf fmt " div_gt%i\n" i;
+ fprintf fmt " (fun n x y => div_gt%i x (GenBase.get_low %s (S n) y))\n" i (pz i);
+ fprintf fmt " w%i_divn1\n" i;
+ done;
+ fprintf fmt " div_gtnm).\n";
+ fprintf fmt "\n";
+
+ fprintf fmt " Theorem spec_div_gt: forall x y,\n";
+ fprintf fmt " [x] > [y] -> 0 < [y] ->\n";
+ fprintf fmt " let (q,r) := div_gt x y in\n";
+ fprintf fmt " [q] = [x] / [y] /\\ [r] = [x] mod [y].\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " assert (FO:\n";
+ fprintf fmt " forall x y, [x] > [y] -> 0 < [y] ->\n";
+ fprintf fmt " let (q,r) := div_gt x y in\n";
+ fprintf fmt " [x] = [q] * [y] + [r] /\\ 0 <= [r] < [y]).\n";
+ fprintf fmt " refine (spec_iter (t_*t_) (fun x y res => x > y -> 0 < y ->\n"; fprintf fmt " let (q,r) := res in\n";
+ fprintf fmt " x = [q] * y + [r] /\\ 0 <= [r] < y)\n";
+ for i = 0 to size do
+ fprintf fmt " div_gt%i\n" i;
+ fprintf fmt " (fun n x y => div_gt%i x (GenBase.get_low %s (S n) y))\n" i (pz i);
+ fprintf fmt " w%i_divn1 _ _ _\n" i;
+ done;
+ fprintf fmt " div_gtnm _).\n";
+ for i = 0 to size do
+ fprintf fmt " intros x y H1 H2; unfold div_gt%i, w%i_div_gt.\n" i i;
+ fprintf fmt " generalize (spec_div_gt w%i_spec x y H1 H2); case znz_div_gt.\n" i;
+ fprintf fmt " intros xx yy; repeat rewrite spec_reduce_%i; auto.\n" i;
+ if i == size then
+ fprintf fmt " intros n x y H2 H3; unfold div_gt%i, w%i_div_gt.\n" i i
+ else
+ fprintf fmt " intros n x y H1 H2 H3; unfold div_gt%i, w%i_div_gt.\n" i i;
+ fprintf fmt " generalize (spec_div_gt w%i_spec x \n" i;
+ fprintf fmt " (GenBase.get_low %s (S n) y)).\n" (pz i);
+ fprintf fmt " ";
+ for j = 0 to i do
+ fprintf fmt "unfold w%i;" (i-j);
+ done;
+ fprintf fmt "case znz_div_gt.\n";
+ fprintf fmt " intros xx yy H4; repeat rewrite spec_reduce_%i.\n" i;
+ fprintf fmt " generalize (spec_get_end%i (S n) y x); unfold to_Z; intros H5.\n" i;
+ fprintf fmt " unfold to_Z in H2; rewrite H5 in H4; auto with zarith.\n";
+ if i == size then
+ fprintf fmt " intros n x y H2 H3.\n"
+ else
+ fprintf fmt " intros n x y H1 H2 H3.\n";
+ fprintf fmt " generalize\n";
+ fprintf fmt " (spec_divn1 w%i w%i_op w%i_spec (S n) x y H3).\n" i i i;
+ fprintf fmt " unfold w%i_divn1;" i;
+ for j = 0 to i do
+ fprintf fmt "unfold w%i;" (i-j);
+ done;
+ fprintf fmt " case gen_divn1.\n";
+ fprintf fmt " intros xx yy H4.\n";
+ if i == size then
+ begin
+ fprintf fmt " repeat rewrite <- spec_gen_eval%in in H4; auto.\n" i;
+ fprintf fmt " rewrite spec_eval%in; auto.\n" i;
+ end
+ else
+ begin
+ fprintf fmt " rewrite to_Z%i_spec; auto with zarith.\n" i;
+ fprintf fmt " repeat rewrite <- spec_gen_eval%in in H4; auto.\n" i;
+ end;
+ done;
+ fprintf fmt " intros n m x y H1 H2; unfold div_gtnm.\n";
+ fprintf fmt " generalize (spec_div_gt (wn_spec (Max.max n m))\n";
+ fprintf fmt " (castm (diff_r n m)\n";
+ fprintf fmt " (extend_tr x (snd (diff n m))))\n";
+ fprintf fmt " (castm (diff_l n m)\n";
+ fprintf fmt " (extend_tr y (fst (diff n m))))).\n";
+ fprintf fmt " case znz_div_gt.\n";
+ fprintf fmt " intros xx yy HH.\n";
+ fprintf fmt " repeat rewrite spec_reduce_n.\n";
+ fprintf fmt " rewrite <- (spec_cast_l n m x).\n";
+ fprintf fmt " rewrite <- (spec_cast_r n m y).\n";
+ fprintf fmt " unfold to_Z; apply HH.\n";
+ fprintf fmt " rewrite <- (spec_cast_l n m x) in H1; auto.\n";
+ fprintf fmt " rewrite <- (spec_cast_r n m y) in H1; auto.\n";
+ fprintf fmt " rewrite <- (spec_cast_r n m y) in H2; auto.\n";
+ fprintf fmt " intros x y H1 H2; generalize (FO x y H1 H2); case div_gt.\n";
+ fprintf fmt " intros q r (H3, H4); split.\n";
+ fprintf fmt " apply (ZDivModAux.Zdiv_unique [x] [y] [q] [r]); auto.\n";
+ fprintf fmt " rewrite Zmult_comm; auto.\n";
+ fprintf fmt " apply (ZDivModAux.Zmod_unique [x] [y] [q] [r]); auto.\n";
+ fprintf fmt " rewrite Zmult_comm; auto.\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
fprintf fmt "\n";
fprintf fmt " Definition div_eucl x y :=\n";
@@ -670,70 +2036,150 @@ let print_Make () =
fprintf fmt " end.\n";
fprintf fmt "\n";
+ fprintf fmt " Theorem spec_div_eucl: forall x y,\n";
+ fprintf fmt " 0 < [y] ->\n";
+ fprintf fmt " let (q,r) := div_eucl x y in\n";
+ fprintf fmt " [q] = [x] / [y] /\\ [r] = [x] mod [y].\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " assert (F0: [zero] = 0).\n";
+ fprintf fmt " exact (spec_0 w0_spec).\n";
+ fprintf fmt " assert (F1: [one] = 1).\n";
+ fprintf fmt " exact (spec_1 w0_spec).\n";
+ fprintf fmt " intros x y H; generalize (spec_compare x y);\n";
+ fprintf fmt " unfold div_eucl; case compare; try rewrite F0;\n";
+ fprintf fmt " try rewrite F1; intros; try split; auto with zarith.\n";
+ fprintf fmt " rewrite H0; apply sym_equal; apply Z_div_same; auto with zarith.\n";
+ fprintf fmt " rewrite H0; apply sym_equal; apply Z_mod_same; auto with zarith.\n";
+ fprintf fmt " apply sym_equal; apply ZDivModAux.Zdiv_lt_0.\n";
+ fprintf fmt " generalize (to_Z_pos x); auto with zarith.\n";
+ fprintf fmt " apply sym_equal; apply ZAux.Zmod_def_small; auto with zarith.\n";
+ fprintf fmt " generalize (to_Z_pos x); auto with zarith.\n";
+ fprintf fmt " apply (spec_div_gt x y); auto.\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+
fprintf fmt " Definition div x y := fst (div_eucl x y).\n";
fprintf fmt "\n";
-
- (* Modulo *)
+
+ fprintf fmt " Theorem spec_div:\n";
+ fprintf fmt " forall x y, 0 < [y] -> [div x y] = [x] / [y].\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros x y H1; unfold div; generalize (spec_div_eucl x y H1);\n";
+ fprintf fmt " case div_eucl; simpl fst.\n";
+ fprintf fmt " intros xx yy (H2, H3); auto with zarith.\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+
+ fprintf fmt " (***************************************************************)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (* Modulo *)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (***************************************************************)\n\n";
+
for i = 0 to size do
fprintf fmt " Definition w%i_mod_gt := w%i_op.(znz_mod_gt).\n" i i
done;
fprintf fmt "\n";
-
+
for i = 0 to size do
- fprintf fmt " Definition w%i_modn1 :=\n" i;
+ fprintf fmt " Definition w%i_modn1 :=\n" i;
fprintf fmt " gen_modn1 w%i_op.(znz_zdigits) w%i_op.(znz_0)\n" i i;
- fprintf fmt
- " w%i_op.(znz_head0) w%i_op.(znz_add_mul_div) w%i_op.(znz_div21)\n"
- i i i;
- fprintf fmt
- " w%i_op.(znz_compare) w%i_op.(znz_sub).\n"
- i i;
+ fprintf fmt " w%i_op.(znz_head0) w%i_op.(znz_add_mul_div) w%i_op.(znz_div21)\n" i i i;
+ fprintf fmt " w%i_op.(znz_compare) w%i_op.(znz_sub).\n" i i;
done;
fprintf fmt "\n";
- fprintf fmt " Definition mod_gt x y :=\n";
- fprintf fmt " match x, y with\n";
- for i = 0 to size do
- for j = 0 to size do
- fprintf fmt " | %s%i wx, %s%i wy =>"
- c i c j;
- if i = j then
- fprintf fmt " reduce_%i (w%i_mod_gt wx wy)\n" i i
- else if i > j then
- fprintf fmt
- " reduce_%i (w%i_modn1 %i wx wy)\n" j j (i-j)
- else begin (* i < j *)
- fprintf fmt
- "\n let wx':= GenBase.extend w%i_0W %i wx in\n"
- i (j-i-1);
- fprintf fmt " reduce_%i (w%i_mod_gt wx' wy)\n" j j;
- end
- done;
- fprintf fmt " | %s%i wx, %sn n wy =>\n" c i c;
- fprintf fmt
- " let wx':= extend n w%i (GenBase.extend w%i_0W %i wx) in\n"
- size i (size-i);
- fprintf fmt " reduce_n n ((make_op n).(znz_mod_gt) wx' wy)\n";
- done;
- for j = 0 to size do
- fprintf fmt " | %sn n wx, %s%i wy =>\n" c c j;
- if j < size then
- fprintf fmt " let wy':= GenBase.extend w%i_0W %i wy in\n"
- j (size-j-1)
- else
- fprintf fmt " let wy':= wy in\n";
- fprintf fmt " reduce_%i (w%i_modn1 (S n) wx wy')\n" size size;
- done;
- fprintf fmt " | %sn n wx, %sn m wy =>\n" c c;
+ fprintf fmt " Let mod_gtnm n m wx wy :=\n";
fprintf fmt " let mn := Max.max n m in\n";
fprintf fmt " let d := diff n m in\n";
fprintf fmt " let op := make_op mn in\n";
- fprintf fmt " reduce_n mn (op.(znz_mod_gt)\n";
- fprintf fmt " (castm (diff_r n m) (extend_tr wx (snd d)))\n";
- fprintf fmt " (castm (diff_l n m) (extend_tr wy (fst d))))\n";
- fprintf fmt " end.\n";
+ fprintf fmt " reduce_n mn (op.(znz_mod_gt)\n";
+ fprintf fmt " (castm (diff_r n m) (extend_tr wx (snd d)))\n";
+ fprintf fmt " (castm (diff_l n m) (extend_tr wy (fst d)))).\n";
fprintf fmt "\n";
-
+
+ fprintf fmt " Definition mod_gt := Eval lazy beta delta[iter] in\n";
+ fprintf fmt " (iter _ \n";
+ for i = 0 to size do
+ fprintf fmt " (fun x y => reduce_%i (w%i_mod_gt x y))\n" i i;
+ fprintf fmt " (fun n x y => reduce_%i (w%i_mod_gt x (GenBase.get_low %s (S n) y)))\n" i i (pz i);
+ fprintf fmt " (fun n x y => reduce_%i (w%i_modn1 (S n) x y))\n" i i;
+ done;
+ fprintf fmt " mod_gtnm).\n";
+ fprintf fmt "\n";
+
+ if gen_proof then
+ begin
+ fprintf fmt " Let spec_modn1 ww (ww_op: znz_op ww) (ww_spec: znz_spec ww_op) := \n";
+ fprintf fmt " (spec_gen_modn1 \n";
+ fprintf fmt " ww_op.(znz_zdigits) ww_op.(znz_0)\n";
+ fprintf fmt " ww_op.(znz_WW) ww_op.(znz_head0)\n";
+ fprintf fmt " ww_op.(znz_add_mul_div) ww_op.(znz_div21)\n";
+ fprintf fmt " ww_op.(znz_compare) ww_op.(znz_sub) (znz_to_Z ww_op)\n";
+ fprintf fmt " (spec_to_Z ww_spec) \n";
+ fprintf fmt " (spec_zdigits ww_spec)\n";
+ fprintf fmt " (spec_0 ww_spec) (spec_WW ww_spec) (spec_head0 ww_spec)\n";
+ fprintf fmt " (spec_add_mul_div ww_spec) (spec_div21 ww_spec) \n";
+ fprintf fmt " (ZnZ.spec_compare ww_spec) (ZnZ.spec_sub ww_spec)).\n";
+ fprintf fmt "\n";
+ end;
+
+ fprintf fmt " Theorem spec_mod_gt:\n";
+ fprintf fmt " forall x y, [x] > [y] -> 0 < [y] -> [mod_gt x y] = [x] mod [y].\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " refine (spec_iter _ (fun x y res => x > y -> 0 < y ->\n";
+ fprintf fmt " [res] = x mod y)\n";
+ for i = 0 to size do
+ fprintf fmt " (fun x y => reduce_%i (w%i_mod_gt x y))\n" i i;
+ fprintf fmt " (fun n x y => reduce_%i (w%i_mod_gt x (GenBase.get_low %s (S n) y)))\n" i i (pz i);
+ fprintf fmt " (fun n x y => reduce_%i (w%i_modn1 (S n) x y)) _ _ _\n" i i;
+ done;
+ fprintf fmt " mod_gtnm _).\n";
+ for i = 0 to size do
+ fprintf fmt " intros x y H1 H2; rewrite spec_reduce_%i.\n" i;
+ fprintf fmt " exact (spec_mod_gt w%i_spec x y H1 H2).\n" i;
+ if i == size then
+ fprintf fmt " intros n x y H2 H3; rewrite spec_reduce_%i.\n" i
+ else
+ fprintf fmt " intros n x y H1 H2 H3; rewrite spec_reduce_%i.\n" i;
+ fprintf fmt " unfold w%i_mod_gt.\n" i;
+ fprintf fmt " rewrite <- (spec_get_end%i (S n) y x); auto with zarith.\n" i;
+ fprintf fmt " unfold to_Z; apply (spec_mod_gt w%i_spec); auto.\n" i;
+ fprintf fmt " rewrite <- (spec_get_end%i (S n) y x) in H2; auto with zarith.\n" i;
+ fprintf fmt " rewrite <- (spec_get_end%i (S n) y x) in H3; auto with zarith.\n" i;
+ if i == size then
+ fprintf fmt " intros n x y H2 H3; rewrite spec_reduce_%i.\n" i
+ else
+ fprintf fmt " intros n x y H1 H2 H3; rewrite spec_reduce_%i.\n" i;
+ fprintf fmt " unfold w%i_modn1, to_Z; rewrite spec_gen_eval%in.\n" i i;
+ fprintf fmt " apply (spec_modn1 _ _ w%i_spec); auto.\n" i;
+ done;
+ fprintf fmt " intros n m x y H1 H2; unfold mod_gtnm.\n";
+ fprintf fmt " repeat rewrite spec_reduce_n.\n";
+ fprintf fmt " rewrite <- (spec_cast_l n m x).\n";
+ fprintf fmt " rewrite <- (spec_cast_r n m y).\n";
+ fprintf fmt " unfold to_Z; apply (spec_mod_gt (wn_spec (Max.max n m))).\n";
+ fprintf fmt " rewrite <- (spec_cast_l n m x) in H1; auto.\n";
+ fprintf fmt " rewrite <- (spec_cast_r n m y) in H1; auto.\n";
+ fprintf fmt " rewrite <- (spec_cast_r n m y) in H2; auto.\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+
fprintf fmt " Definition modulo x y := \n";
fprintf fmt " match compare x y with\n";
fprintf fmt " | Eq => zero\n";
@@ -742,19 +2188,65 @@ let print_Make () =
fprintf fmt " end.\n";
fprintf fmt "\n";
- (* Definition du gcd *)
+ fprintf fmt " Theorem spec_modulo:\n";
+ fprintf fmt " forall x y, 0 < [y] -> [modulo x y] = [x] mod [y].\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " assert (F0: [zero] = 0).\n";
+ fprintf fmt " exact (spec_0 w0_spec).\n";
+ fprintf fmt " assert (F1: [one] = 1).\n";
+ fprintf fmt " exact (spec_1 w0_spec).\n";
+ fprintf fmt " intros x y H; generalize (spec_compare x y);\n";
+ fprintf fmt " unfold modulo; case compare; try rewrite F0;\n";
+ fprintf fmt " try rewrite F1; intros; try split; auto with zarith.\n";
+ fprintf fmt " rewrite H0; apply sym_equal; apply Z_mod_same; auto with zarith.\n";
+ fprintf fmt " apply sym_equal; apply ZAux.Zmod_def_small; auto with zarith.\n";
+ fprintf fmt " generalize (to_Z_pos x); auto with zarith.\n";
+ fprintf fmt " apply spec_mod_gt; auto.\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+
+ fprintf fmt " (***************************************************************)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (* Gcd *)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (***************************************************************)\n\n";
+
fprintf fmt " Definition digits x :=\n";
fprintf fmt " match x with\n";
- for i = 0 to size do
+ for i = 0 to size do
fprintf fmt " | %s%i _ => w%i_op.(znz_digits)\n" c i i;
done;
fprintf fmt " | %sn n _ => (make_op n).(znz_digits)\n" c;
fprintf fmt " end.\n";
fprintf fmt "\n";
+
+ fprintf fmt " Theorem spec_digits: forall x, 0 <= [x] < 2 ^ Zpos (digits x).\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros x; case x; clear x.\n";
+ for i = 0 to size do
+ fprintf fmt " intros x; unfold to_Z, digits;\n";
+ fprintf fmt " generalize (spec_to_Z w%i_spec x); unfold base; intros H; exact H.\n" i;
+ done;
+ fprintf fmt " intros n x; unfold to_Z, digits;\n";
+ fprintf fmt " generalize (spec_to_Z (wn_spec n) x); unfold base; intros H; exact H.\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+
+
fprintf fmt " Definition gcd_gt_body a b cont :=\n";
fprintf fmt " match compare b zero with\n";
- fprintf fmt " | Gt =>\n";
+ fprintf fmt " | Gt =>\n";
fprintf fmt " let r := mod_gt a b in\n";
fprintf fmt " match compare r zero with\n";
fprintf fmt " | Gt => cont r (mod_gt b r)\n";
@@ -764,15 +2256,118 @@ let print_Make () =
fprintf fmt " end.\n";
fprintf fmt "\n";
- fprintf fmt " Fixpoint gcd_gt (p:positive) (cont:%s->%s->%s) (a b:%s) {struct p} : %s :=\n" t t t t t;
+ if gen_proof then
+ begin
+ fprintf fmt " Theorem Zspec_gcd_gt_body: forall a b cont p,\n";
+ fprintf fmt " [a] > [b] -> [a] < 2 ^ p ->\n";
+ fprintf fmt " (forall a1 b1, [a1] < 2 ^ (p - 1) -> [a1] > [b1] ->\n";
+ fprintf fmt " Zis_gcd [a1] [b1] [cont a1 b1]) -> \n";
+ fprintf fmt " Zis_gcd [a] [b] [gcd_gt_body a b cont].\n";
+ fprintf fmt " Proof.\n";
+ fprintf fmt " assert (F1: [zero] = 0).\n";
+ fprintf fmt " unfold zero, w_0, to_Z; rewrite (spec_0 w0_spec); auto.\n";
+ fprintf fmt " intros a b cont p H2 H3 H4; unfold gcd_gt_body.\n";
+ fprintf fmt " generalize (spec_compare b zero); case compare; try rewrite F1.\n";
+ fprintf fmt " intros HH; rewrite HH; apply Zis_gcd_0.\n";
+ fprintf fmt " intros HH; absurd (0 <= [b]); auto with zarith.\n";
+ fprintf fmt " case (spec_digits b); auto with zarith.\n";
+ fprintf fmt " intros H5; generalize (spec_compare (mod_gt a b) zero); \n";
+ fprintf fmt " case compare; try rewrite F1.\n";
+ fprintf fmt " intros H6; rewrite <- (Zmult_1_r [b]).\n";
+ fprintf fmt " rewrite (Z_div_mod_eq [a] [b]); auto with zarith.\n";
+ fprintf fmt " rewrite <- spec_mod_gt; auto with zarith.\n";
+ fprintf fmt " rewrite H6; rewrite Zplus_0_r.\n";
+ fprintf fmt " apply Zis_gcd_mult; apply Zis_gcd_1.\n";
+ fprintf fmt " intros; apply False_ind.\n";
+ fprintf fmt " case (spec_digits (mod_gt a b)); auto with zarith.\n";
+ fprintf fmt " intros H6; apply GenDiv.Zis_gcd_mod; auto with zarith.\n";
+ fprintf fmt " apply GenDiv.Zis_gcd_mod; auto with zarith.\n";
+ fprintf fmt " rewrite <- spec_mod_gt; auto with zarith.\n";
+ fprintf fmt " assert (F2: [b] > [mod_gt a b]).\n";
+ fprintf fmt " case (Z_mod_lt [a] [b]); auto with zarith.\n";
+ fprintf fmt " repeat rewrite <- spec_mod_gt; auto with zarith.\n";
+ fprintf fmt " assert (F3: [mod_gt a b] > [mod_gt b (mod_gt a b)]).\n";
+ fprintf fmt " case (Z_mod_lt [b] [mod_gt a b]); auto with zarith.\n";
+ fprintf fmt " rewrite <- spec_mod_gt; auto with zarith.\n";
+ fprintf fmt " repeat rewrite <- spec_mod_gt; auto with zarith.\n";
+ fprintf fmt " apply H4; auto with zarith.\n";
+ fprintf fmt " apply Zmult_lt_reg_r with 2; auto with zarith.\n";
+ fprintf fmt " apply Zle_lt_trans with ([b] + [mod_gt a b]); auto with zarith.\n";
+ fprintf fmt " apply Zle_lt_trans with (([a]/[b]) * [b] + [mod_gt a b]); auto with zarith.\n";
+ fprintf fmt " apply Zplus_le_compat_r.\n";
+ fprintf fmt " pattern [b] at 1; rewrite <- (Zmult_1_l [b]).\n";
+ fprintf fmt " apply Zmult_le_compat_r; auto with zarith.\n";
+ fprintf fmt " case (Zle_lt_or_eq 0 ([a]/[b])); auto with zarith.\n";
+ fprintf fmt " intros HH; rewrite (Z_div_mod_eq [a] [b]) in H2;\n";
+ fprintf fmt " try rewrite <- HH in H2; auto with zarith.\n";
+ fprintf fmt " case (Z_mod_lt [a] [b]); auto with zarith.\n";
+ fprintf fmt " rewrite Zmult_comm; rewrite spec_mod_gt; auto with zarith.\n";
+ fprintf fmt " rewrite <- Z_div_mod_eq; auto with zarith.\n";
+ fprintf fmt " pattern 2 at 2; rewrite <- (Zpower_exp_1 2).\n";
+ fprintf fmt " rewrite <- Zpower_exp; auto with zarith.\n";
+ fprintf fmt " ring_simplify (p - 1 + 1); auto.\n";
+ fprintf fmt " case (Zle_lt_or_eq 0 p); auto with zarith.\n";
+ fprintf fmt " generalize H3; case p; simpl Zpower; auto with zarith.\n";
+ fprintf fmt " intros HH; generalize H3; rewrite <- HH; simpl Zpower; auto with zarith.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+ end;
+
+ fprintf fmt " Fixpoint gcd_gt_aux (p:positive) (cont:t->t->t) (a b:t) {struct p} : t :=\n";
fprintf fmt " gcd_gt_body a b\n";
fprintf fmt " (fun a b =>\n";
fprintf fmt " match p with\n";
fprintf fmt " | xH => cont a b\n";
- fprintf fmt " | xO p => gcd_gt p (gcd_gt p cont) a b\n";
- fprintf fmt " | xI p => gcd_gt p (gcd_gt p cont) a b\n";
+ fprintf fmt " | xO p => gcd_gt_aux p (gcd_gt_aux p cont) a b\n";
+ fprintf fmt " | xI p => gcd_gt_aux p (gcd_gt_aux p cont) a b\n";
fprintf fmt " end).\n";
fprintf fmt "\n";
+
+ if gen_proof then
+ begin
+ fprintf fmt " Theorem Zspec_gcd_gt_aux: forall p n a b cont,\n";
+ fprintf fmt " [a] > [b] -> [a] < 2 ^ (Zpos p + n) ->\n";
+ fprintf fmt " (forall a1 b1, [a1] < 2 ^ n -> [a1] > [b1] ->\n";
+ fprintf fmt " Zis_gcd [a1] [b1] [cont a1 b1]) ->\n";
+ fprintf fmt " Zis_gcd [a] [b] [gcd_gt_aux p cont a b].\n";
+ fprintf fmt " intros p; elim p; clear p.\n";
+ fprintf fmt " intros p Hrec n a b cont H2 H3 H4.\n";
+ fprintf fmt " unfold gcd_gt_aux; apply Zspec_gcd_gt_body with (Zpos (xI p) + n); auto.\n";
+ fprintf fmt " intros a1 b1 H6 H7.\n";
+ fprintf fmt " apply Hrec with (Zpos p + n); auto.\n";
+ fprintf fmt " replace (Zpos p + (Zpos p + n)) with\n";
+ fprintf fmt " (Zpos (xI p) + n - 1); auto.\n";
+ fprintf fmt " rewrite Zpos_xI; ring.\n";
+ fprintf fmt " intros a2 b2 H9 H10.\n";
+ fprintf fmt " apply Hrec with n; auto.\n";
+ fprintf fmt " intros p Hrec n a b cont H2 H3 H4.\n";
+ fprintf fmt " unfold gcd_gt_aux; apply Zspec_gcd_gt_body with (Zpos (xO p) + n); auto.\n";
+ fprintf fmt " intros a1 b1 H6 H7.\n";
+ fprintf fmt " apply Hrec with (Zpos p + n - 1); auto.\n";
+ fprintf fmt " replace (Zpos p + (Zpos p + n - 1)) with\n";
+ fprintf fmt " (Zpos (xO p) + n - 1); auto.\n";
+ fprintf fmt " rewrite Zpos_xO; ring.\n";
+ fprintf fmt " intros a2 b2 H9 H10.\n";
+ fprintf fmt " apply Hrec with (n - 1); auto.\n";
+ fprintf fmt " replace (Zpos p + (n - 1)) with\n";
+ fprintf fmt " (Zpos p + n - 1); auto with zarith.\n";
+ fprintf fmt " intros a3 b3 H12 H13; apply H4; auto with zarith.\n";
+ fprintf fmt " apply Zlt_le_trans with (1 := H12).\n";
+ fprintf fmt " case (Zle_or_lt 1 n); intros HH.\n";
+ fprintf fmt " apply Zpower_le_monotone; auto with zarith.\n";
+ fprintf fmt " apply Zle_trans with 0; auto with zarith.\n";
+ fprintf fmt " assert (HH1: n - 1 < 0); auto with zarith.\n";
+ fprintf fmt " generalize HH1; case (n - 1); auto with zarith.\n";
+ fprintf fmt " intros p1 HH2; discriminate.\n";
+ fprintf fmt " intros n a b cont H H2 H3.\n";
+ fprintf fmt " simpl gcd_gt_aux.\n";
+ fprintf fmt " apply Zspec_gcd_gt_body with (n + 1); auto with zarith.\n";
+ fprintf fmt " rewrite Zplus_comm; auto.\n";
+ fprintf fmt " intros a1 b1 H5 H6; apply H3; auto.\n";
+ fprintf fmt " replace n with (n + 1 - 1); auto; try ring.\n";
+ fprintf fmt " Qed.\n";
+ fprintf fmt "\n";
+ end;
fprintf fmt " Definition gcd_cont a b :=\n";
fprintf fmt " match compare one b with\n";
@@ -781,34 +2376,156 @@ let print_Make () =
fprintf fmt " end.\n";
fprintf fmt "\n";
+ fprintf fmt " Definition gcd_gt a b := gcd_gt_aux (digits a) gcd_cont a b.\n";
+ fprintf fmt "\n";
+
+ fprintf fmt " Theorem spec_gcd_gt: forall a b,\n";
+ fprintf fmt " [a] > [b] -> [gcd_gt a b] = Zgcd [a] [b].\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros a b H2.\n";
+ fprintf fmt " case (spec_digits (gcd_gt a b)); intros H3 H4.\n";
+ fprintf fmt " case (spec_digits a); intros H5 H6.\n";
+ fprintf fmt " apply sym_equal; apply Zis_gcd_gcd; auto with zarith.\n";
+ fprintf fmt " unfold gcd_gt; apply Zspec_gcd_gt_aux with 0; auto with zarith.\n";
+ fprintf fmt " intros a1 a2; rewrite Zpower_exp_0.\n";
+ fprintf fmt " case (spec_digits a2); intros H7 H8;\n";
+ fprintf fmt " intros; apply False_ind; auto with zarith.\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+
fprintf fmt " Definition gcd a b :=\n";
fprintf fmt " match compare a b with\n";
fprintf fmt " | Eq => a\n";
- fprintf fmt " | Lt => gcd_gt (digits b) gcd_cont b a\n";
- fprintf fmt " | Gt => gcd_gt (digits a) gcd_cont a b\n";
+ fprintf fmt " | Lt => gcd_gt b a\n";
+ fprintf fmt " | Gt => gcd_gt a b\n";
fprintf fmt " end.\n";
fprintf fmt "\n";
- fprintf fmt "Definition pheight p := Peano.pred (nat_of_P (get_height w0_op.(znz_digits) (plength p))).";
+ fprintf fmt " Theorem spec_gcd: forall a b, [gcd a b] = Zgcd [a] [b].\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros a b.\n";
+ fprintf fmt " case (spec_digits a); intros H1 H2.\n";
+ fprintf fmt " case (spec_digits b); intros H3 H4.\n";
+ fprintf fmt " unfold gcd; generalize (spec_compare a b); case compare.\n";
+ fprintf fmt " intros HH; rewrite HH; apply sym_equal; apply Zis_gcd_gcd; auto.\n";
+ fprintf fmt " apply Zis_gcd_refl.\n";
+ fprintf fmt " intros; apply trans_equal with (Zgcd [b] [a]).\n";
+ fprintf fmt " apply spec_gcd_gt; auto with zarith.\n";
+ fprintf fmt " apply Zis_gcd_gcd; auto with zarith.\n";
+ fprintf fmt " apply Zgcd_is_pos.\n";
+ fprintf fmt " apply Zis_gcd_sym; apply Zgcd_is_gcd.\n";
+ fprintf fmt " intros; apply spec_gcd_gt; auto.\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+
+
+ fprintf fmt " (***************************************************************)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (* Conversion *)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (***************************************************************)\n\n";
+
+ fprintf fmt " Definition pheight p := \n";
+ fprintf fmt " Peano.pred (nat_of_P (get_height w0_op.(znz_digits) (plength p))).\n";
+ fprintf fmt "\n";
+
+ fprintf fmt " Theorem pheight_correct: forall p, \n";
+ fprintf fmt " Zpos p < 2 ^ (Zpos (znz_digits w0_op) * 2 ^ (Z_of_nat (pheight p))).\n";
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros p; unfold pheight.\n";
+ fprintf fmt " assert (F1: forall x, Z_of_nat (Peano.pred (nat_of_P x)) = Zpos x - 1).\n";
+ fprintf fmt " intros x.\n";
+ fprintf fmt " assert (Zsucc (Z_of_nat (Peano.pred (nat_of_P x))) = Zpos x); auto with zarith.\n";
+ fprintf fmt " rewrite <- inj_S.\n";
+ fprintf fmt " rewrite <- (fun x => S_pred x 0); auto with zarith.\n";
+ fprintf fmt " rewrite Zpos_eq_Z_of_nat_o_nat_of_P; auto.\n";
+ fprintf fmt " apply lt_le_trans with 1%snat; auto with zarith.\n" "%";
+ fprintf fmt " exact (le_Pmult_nat x 1).\n";
+ fprintf fmt " rewrite F1; clear F1.\n";
+ fprintf fmt " assert (F2:= (get_height_correct (znz_digits w0_op) (plength p))).\n";
+ fprintf fmt " apply Zlt_le_trans with (Zpos (Psucc p)).\n";
+ fprintf fmt " rewrite Zpos_succ_morphism; auto with zarith.\n";
+ fprintf fmt " apply Zle_trans with (1 := plength_pred_correct (Psucc p)).\n";
+ fprintf fmt " rewrite Ppred_succ.\n";
+ fprintf fmt " apply Zpower_le_monotone; auto with zarith.\n";
+ fprintf fmt " Qed.\n";
fprintf fmt "\n";
fprintf fmt " Definition of_pos x :=\n";
- fprintf fmt " let h := pheight x in\n";
+ fprintf fmt " let h := pheight x in\n";
fprintf fmt " match h with\n";
- let rec print_S s fmt i =
- if i = 0 then fprintf fmt "%s" s
- else fprintf fmt "(S %a)" (print_S s) (i-1)
- in
for i = 0 to size do
- fprintf fmt " | ";
- print_S "O" fmt i;
- fprintf fmt " => %s%i (snd (w%i_op.(znz_of_pos) x))\n" "reduce_" i i
+ fprintf fmt " | %i%snat => reduce_%i (snd (w%i_op.(znz_of_pos) x))\n" i "%" i i;
done;
fprintf fmt " | _ =>\n";
- fprintf fmt " let n := minus h %i in\n" (size+1);
- fprintf fmt " %sn n (snd ((make_op n).(znz_of_pos) x))\n" "reduce_";
+ fprintf fmt " let n := minus h %i in\n" (size + 1);
+ fprintf fmt " reduce_n n (snd ((make_op n).(znz_of_pos) x))\n";
fprintf fmt " end.\n";
fprintf fmt "\n";
+
+ fprintf fmt " Theorem spec_of_pos: forall x,\n";
+ fprintf fmt " [of_pos x] = Zpos x.\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " assert (F := spec_more_than_1_digit w0_spec).\n";
+ fprintf fmt " intros x; unfold of_pos; case_eq (pheight x).\n";
+ for i = 0 to size do
+ if i <> 0 then
+ fprintf fmt " intros n; case n; clear n.\n";
+ fprintf fmt " intros H1; rewrite spec_reduce_%i; unfold to_Z.\n" i;
+ fprintf fmt " apply (znz_of_pos_correct w%i_spec).\n" i;
+ fprintf fmt " apply Zlt_le_trans with (1 := pheight_correct x).\n";
+ fprintf fmt " rewrite H1; simpl Z_of_nat; change (2^%i) with (%s).\n" i (gen2 i);
+ fprintf fmt " unfold base.\n";
+ fprintf fmt " apply Zpower_le_monotone; split; auto with zarith.\n";
+ if i <> 0 then
+ begin
+ fprintf fmt " rewrite Zmult_comm; repeat rewrite <- Zmult_assoc.\n";
+ fprintf fmt " repeat rewrite <- Zpos_xO.\n";
+ fprintf fmt " refine (Zle_refl _).\n";
+ end;
+ done;
+ fprintf fmt " intros n.\n";
+ fprintf fmt " intros H1; rewrite spec_reduce_n; unfold to_Z.\n";
+ fprintf fmt " simpl minus; rewrite <- minus_n_O.\n";
+ fprintf fmt " apply (znz_of_pos_correct (wn_spec n)).\n";
+ fprintf fmt " apply Zlt_le_trans with (1 := pheight_correct x).\n";
+ fprintf fmt " unfold base.\n";
+ fprintf fmt " apply Zpower_le_monotone; auto with zarith.\n";
+ fprintf fmt " split; auto with zarith.\n";
+ fprintf fmt " rewrite H1.\n";
+ fprintf fmt " elim n; clear n H1.\n";
+ fprintf fmt " simpl Z_of_nat; change (2^%i) with (%s).\n" (size + 1) (gen2 (size + 1));
+ fprintf fmt " rewrite Zmult_comm; repeat rewrite <- Zmult_assoc.\n";
+ fprintf fmt " repeat rewrite <- Zpos_xO.\n";
+ fprintf fmt " refine (Zle_refl _).\n";
+ fprintf fmt " intros n Hrec.\n";
+ fprintf fmt " rewrite make_op_S.\n";
+ fprintf fmt " change (%sznz_digits (word _ (S (S n))) (mk_zn2z_op_karatsuba (make_op n))) with\n" "@";
+ fprintf fmt " (xO (znz_digits (make_op n))).\n";
+ fprintf fmt " rewrite (fun x y => (Zpos_xO (%sznz_digits x y))).\n" "@";
+ fprintf fmt " rewrite inj_S; unfold Zsucc.\n";
+ fprintf fmt " rewrite Zplus_comm; rewrite Zpower_exp; auto with zarith.\n";
+ fprintf fmt " rewrite Zpower_exp_1.\n";
+ fprintf fmt " assert (tmp: forall x y z, x * (y * z) = y * (x * z));\n";
+ fprintf fmt " [intros; ring | rewrite tmp; clear tmp].\n";
+ fprintf fmt " apply Zmult_le_compat_l; auto with zarith.\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
fprintf fmt " Definition of_N x :=\n";
fprintf fmt " match x with\n";
@@ -817,15 +2534,29 @@ let print_Make () =
fprintf fmt " end.\n";
fprintf fmt "\n";
- fprintf fmt " Definition to_Z x :=\n";
- fprintf fmt " match x with\n";
- for i = 0 to size do
- fprintf fmt " | %s%i wx => w%i_op.(znz_to_Z) wx\n" c i i
- done;
- fprintf fmt " | %sn n wx => (make_op n).(znz_to_Z) wx\n" c;
- fprintf fmt " end.\n";
+
+ fprintf fmt " Theorem spec_of_N: forall x,\n";
+ fprintf fmt " [of_N x] = Z_of_N x.\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros x; case x.\n";
+ fprintf fmt " simpl of_N.\n";
+ fprintf fmt " unfold zero, w_0, to_Z; rewrite (spec_0 w0_spec); auto.\n";
+ fprintf fmt " intros p; exact (spec_of_pos p).\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
fprintf fmt "\n";
+ fprintf fmt " (***************************************************************)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (* Shift *)\n";
+ fprintf fmt " (* *)\n";
+ fprintf fmt " (***************************************************************)\n\n";
+
+
(* Head0 *)
fprintf fmt " Definition head0 w := match w with\n";
@@ -836,6 +2567,54 @@ let print_Make () =
fprintf fmt " end.\n";
fprintf fmt "\n";
+ fprintf fmt " Theorem spec_head00: forall x, [x] = 0 ->[head0 x] = Zpos (digits x).\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros x; case x; unfold head0; clear x.\n";
+ for i = 0 to size do
+ fprintf fmt " intros x; rewrite spec_reduce_%i; exact (spec_head00 w%i_spec x).\n" i i;
+ done;
+ fprintf fmt " intros n x; rewrite spec_reduce_n; exact (spec_head00 (wn_spec n) x).\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt " \n";
+
+ fprintf fmt " Theorem spec_head0: forall x, 0 < [x] ->\n";
+ fprintf fmt " 2 ^ (Zpos (digits x) - 1) <= 2 ^ [head0 x] * [x] < 2 ^ Zpos (digits x).\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " assert (F0: forall x, (x - 1) + 1 = x).\n";
+ fprintf fmt " intros; ring. \n";
+ fprintf fmt " intros x; case x; unfold digits, head0; clear x.\n";
+ for i = 0 to size do
+ fprintf fmt " intros x Hx; rewrite spec_reduce_%i.\n" i;
+ fprintf fmt " assert (F1:= spec_more_than_1_digit w%i_spec).\n" i;
+ fprintf fmt " generalize (spec_head0 w%i_spec x Hx).\n" i;
+ fprintf fmt " unfold base.\n";
+ fprintf fmt " pattern (Zpos (znz_digits w%i_op)) at 1; \n" i;
+ fprintf fmt " rewrite <- (fun x => (F0 (Zpos x))).\n";
+ fprintf fmt " rewrite Zpower_exp; auto with zarith.\n";
+ fprintf fmt " rewrite Zpower_exp_1; rewrite Z_div_mult; auto with zarith.\n";
+ done;
+ fprintf fmt " intros n x Hx; rewrite spec_reduce_n.\n";
+ fprintf fmt " assert (F1:= spec_more_than_1_digit (wn_spec n)).\n";
+ fprintf fmt " generalize (spec_head0 (wn_spec n) x Hx).\n";
+ fprintf fmt " unfold base.\n";
+ fprintf fmt " pattern (Zpos (znz_digits (make_op n))) at 1; \n";
+ fprintf fmt " rewrite <- (fun x => (F0 (Zpos x))).\n";
+ fprintf fmt " rewrite Zpower_exp; auto with zarith.\n";
+ fprintf fmt " rewrite Zpower_exp_1; rewrite Z_div_mult; auto with zarith.\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+
+
(* Tail0 *)
fprintf fmt " Definition tail0 w := match w with\n";
for i = 0 to size do
@@ -845,6 +2624,40 @@ let print_Make () =
fprintf fmt " end.\n";
fprintf fmt "\n";
+
+ fprintf fmt " Theorem spec_tail00: forall x, [x] = 0 ->[tail0 x] = Zpos (digits x).\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros x; case x; unfold tail0; clear x.\n";
+ for i = 0 to size do
+ fprintf fmt " intros x; rewrite spec_reduce_%i; exact (spec_tail00 w%i_spec x).\n" i i;
+ done;
+ fprintf fmt " intros n x; rewrite spec_reduce_n; exact (spec_tail00 (wn_spec n) x).\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt " \n";
+
+
+ fprintf fmt " Theorem spec_tail0: forall x,\n";
+ fprintf fmt " 0 < [x] -> exists y, 0 <= y /\\ [x] = (2 * y + 1) * 2 ^ [tail0 x].\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros x; case x; clear x; unfold tail0.\n";
+ for i = 0 to size do
+ fprintf fmt " intros x Hx; rewrite spec_reduce_%i; exact (spec_tail0 w%i_spec x Hx).\n" i i;
+ done;
+ fprintf fmt " intros n x Hx; rewrite spec_reduce_n; exact (spec_tail0 (wn_spec n) x Hx).\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+
+
(* Number of digits *)
fprintf fmt " Definition %sdigits x :=\n" c;
fprintf fmt " match x with\n";
@@ -852,58 +2665,26 @@ let print_Make () =
for i = 1 to size do
fprintf fmt " | %s%i _ => reduce_%i w%i_op.(znz_zdigits)\n" c i i i;
done;
- fprintf fmt " | %sn n _ => reduce_n n (make_op n).(znz_zdigits)\n \n" c;
+ fprintf fmt " | %sn n _ => reduce_n n (make_op n).(znz_zdigits)\n" c;
fprintf fmt " end.\n";
fprintf fmt "\n";
-
- fprintf fmt " Definition level ";
- for i = 0 to size do
- fprintf fmt "f%i " i;
- done;
- fprintf fmt " fn x y: %s_ := match x, y with\n" t;
- fprintf fmt " | %s0 wx, %s0 wy => f0 wx wy \n" c c;
- for j = 1 to size do
- fprintf fmt " | %s0 wx, %s%i wy => f%i " c c j j;
- if j = 1 then fprintf fmt "(WW w_0 wx) wy\n"
- else fprintf fmt "(extend%i w0 (WW w_0 wx)) wy\n" (j-1)
- done;
- fprintf fmt " | %s0 wx, %sn n wy =>\n" c c;
- fprintf fmt " fn n (extend n w%i (extend%i w0 (WW w_0 wx))) wy\n"
- size size;
- for i = 1 to size do
- fprintf fmt " | %s%i wx, %s0 wy =>\n" c i c;
- fprintf fmt
- " f%i wx " i;
- if i = 1 then fprintf fmt "(WW w_0 wy)\n"
- else fprintf fmt "(extend%i w0 (WW w_0 wy))\n" (i-1);
- for j = 1 to size do
- fprintf fmt " | %s%i wx, %s%i wy => " c i c j;
- if i < j then fprintf fmt "f%i (extend%i w%i wx) wy\n" j (j-i) (i-1)
- else if i = j then fprintf fmt "f%i wx wy\n" j
- else fprintf fmt "f%i wx (extend%i w%i wy)\n" i (i-j) (j-1)
- done;
- fprintf fmt
- " | %s%i wx, %sn n wy => fn n (extend n w%i (extend%i w%i wx)) wy\n"
- c i c size (size-i+1) (i-1)
- done;
- fprintf fmt " | %sn n wx, %s0 wy =>\n" c c;
- fprintf fmt " fn n wx (extend n w%i (extend%i w0 (WW w_0 wy)))\n"
- size size;
- for j = 1 to size do
- fprintf fmt
- " | %sn n wx, %s%i wy => fn n wx (extend n w%i (extend%i w%i wy))\n"
- c c j size (size-j+1) (j-1);
- done;
- fprintf fmt " | %sn n wx, %sn m wy =>\n" c c;
- fprintf fmt " let mn := Max.max n m in\n";
- fprintf fmt " let d := diff n m in\n";
- fprintf fmt " fn mn\n";
- fprintf fmt " (castm (diff_r n m) (extend_tr wx (snd d)))\n";
- fprintf fmt " (castm (diff_l n m) (extend_tr wy (fst d)))\n";
- fprintf fmt " end.\n";
+ fprintf fmt " Theorem spec_Ndigits: forall x, [Ndigits x] = Zpos (digits x).\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros x; case x; clear x; unfold Ndigits, digits.\n";
+ for i = 0 to size do
+ fprintf fmt " intros _; try rewrite spec_reduce_%i; exact (spec_zdigits w%i_spec).\n" i i;
+ done;
+ fprintf fmt " intros n _; try rewrite spec_reduce_n; exact (spec_zdigits (wn_spec n)).\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
fprintf fmt "\n";
+
(* Shiftr *)
for i = 0 to size do
fprintf fmt " Definition shiftr%i n x := w%i_op.(znz_add_mul_div) (w%i_op.(znz_sub) w%i_op.(znz_zdigits) n) w%i_op.(znz_0) x.\n" i i i i i;
@@ -911,9 +2692,8 @@ let print_Make () =
fprintf fmt " Definition shiftrn n p x := (make_op n).(znz_add_mul_div) ((make_op n).(znz_sub) (make_op n).(znz_zdigits) p) (make_op n).(znz_0) x.\n";
fprintf fmt "\n";
- fprintf fmt " Definition shiftr := \n";
- fprintf fmt " Eval lazy beta delta [level] in \n";
- fprintf fmt " level (fun n x => %s0 (shiftr0 n x))\n" c;
+ fprintf fmt " Definition shiftr := Eval lazy beta delta [same_level] in \n";
+ fprintf fmt " same_level _ (fun n x => %s0 (shiftr0 n x))\n" c;
for i = 1 to size do
fprintf fmt " (fun n x => reduce_%i (shiftr%i n x))\n" i i;
done;
@@ -921,6 +2701,147 @@ let print_Make () =
fprintf fmt "\n";
+ fprintf fmt " Theorem spec_shiftr: forall n x,\n";
+ fprintf fmt " [n] <= [Ndigits x] -> [shiftr n x] = [x] / 2 ^ [n].\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " assert (F0: forall x y, x - (x - y) = y).\n";
+ fprintf fmt " intros; ring.\n";
+ fprintf fmt " assert (F2: forall x y z, 0 <= x -> 0 <= y -> x < z -> 0 <= x / 2 ^ y < z).\n";
+ fprintf fmt " intros x y z HH HH1 HH2.\n";
+ fprintf fmt " split; auto with zarith.\n";
+ fprintf fmt " apply Zle_lt_trans with (2 := HH2); auto with zarith.\n";
+ fprintf fmt " apply ZDivModAux.Zdiv_le_upper_bound; auto with zarith.\n";
+ fprintf fmt " pattern x at 1; replace x with (x * 2 ^ 0); auto with zarith.\n";
+ fprintf fmt " apply Zmult_le_compat_l; auto.\n";
+ fprintf fmt " apply Zpower_le_monotone; auto with zarith.\n";
+ fprintf fmt " rewrite Zpower_exp_0; ring.\n";
+ fprintf fmt " assert (F3: forall x y, 0 <= y -> y <= x -> 0 <= x - y < 2 ^ x).\n";
+ fprintf fmt " intros xx y HH HH1.\n";
+ fprintf fmt " split; auto with zarith.\n";
+ fprintf fmt " apply Zle_lt_trans with xx; auto with zarith.\n";
+ fprintf fmt " apply ZPowerAux.Zpower2_lt_lin; auto with zarith.\n";
+ fprintf fmt " assert (F4: forall ww ww1 ww2 \n";
+ fprintf fmt " (ww_op: znz_op ww) (ww1_op: znz_op ww1) (ww2_op: znz_op ww2)\n";
+ fprintf fmt " xx yy xx1 yy1,\n";
+ fprintf fmt " znz_to_Z ww2_op yy <= znz_to_Z ww1_op (znz_zdigits ww1_op) ->\n";
+ fprintf fmt " znz_to_Z ww1_op (znz_zdigits ww1_op) <= znz_to_Z ww_op (znz_zdigits ww_op) ->\n";
+ fprintf fmt " znz_spec ww_op -> znz_spec ww1_op -> znz_spec ww2_op ->\n";
+ fprintf fmt " znz_to_Z ww_op xx1 = znz_to_Z ww1_op xx ->\n";
+ fprintf fmt " znz_to_Z ww_op yy1 = znz_to_Z ww2_op yy ->\n";
+ fprintf fmt " znz_to_Z ww_op\n";
+ fprintf fmt " (znz_add_mul_div ww_op (znz_sub ww_op (znz_zdigits ww_op) yy1)\n";
+ fprintf fmt " (znz_0 ww_op) xx1) = znz_to_Z ww1_op xx / 2 ^ znz_to_Z ww2_op yy).\n";
+ fprintf fmt " intros ww ww1 ww2 ww_op ww1_op ww2_op xx yy xx1 yy1 Hl Hl1 Hw Hw1 Hw2 Hx Hy.\n";
+ fprintf fmt " case (spec_to_Z Hw xx1); auto with zarith; intros HH1 HH2.\n";
+ fprintf fmt " case (spec_to_Z Hw yy1); auto with zarith; intros HH3 HH4.\n";
+ fprintf fmt " rewrite <- Hx.\n";
+ fprintf fmt " rewrite <- Hy.\n";
+ fprintf fmt " generalize (spec_add_mul_div Hw\n";
+ fprintf fmt " (znz_0 ww_op) xx1\n";
+ fprintf fmt " (znz_sub ww_op (znz_zdigits ww_op) \n";
+ fprintf fmt " yy1)\n";
+ fprintf fmt " ).\n";
+ fprintf fmt " rewrite (spec_0 Hw).\n";
+ fprintf fmt " rewrite Zmult_0_l; rewrite Zplus_0_l.\n";
+ fprintf fmt " rewrite (ZnZ.spec_sub Hw).\n";
+ fprintf fmt " rewrite Zmod_def_small; auto with zarith.\n";
+ fprintf fmt " rewrite (spec_zdigits Hw).\n";
+ fprintf fmt " rewrite F0.\n";
+ fprintf fmt " rewrite Zmod_def_small; auto with zarith.\n";
+ fprintf fmt " unfold base; rewrite (spec_zdigits Hw) in Hl1 |- *;\n";
+ fprintf fmt " auto with zarith.\n";
+ fprintf fmt " assert (F5: forall n m, (n <= m)%snat ->\n" "%";
+ fprintf fmt " Zpos (znz_digits (make_op n)) <= Zpos (znz_digits (make_op m))).\n";
+ fprintf fmt " intros n m HH; elim HH; clear m HH; auto with zarith.\n";
+ fprintf fmt " intros m HH Hrec; apply Zle_trans with (1 := Hrec).\n";
+ fprintf fmt " rewrite make_op_S.\n";
+ fprintf fmt " match goal with |- Zpos ?Y <= ?X => change X with (Zpos (xO Y)) end.\n";
+ fprintf fmt " rewrite Zpos_xO.\n";
+ fprintf fmt " assert (0 <= Zpos (znz_digits (make_op n))); auto with zarith.\n";
+ fprintf fmt " assert (F6: forall n, Zpos (znz_digits w%i_op) <= Zpos (znz_digits (make_op n))).\n" size;
+ fprintf fmt " intros n ; apply Zle_trans with (Zpos (znz_digits (make_op 0))).\n";
+ fprintf fmt " change (znz_digits (make_op 0)) with (xO (znz_digits w%i_op)).\n" size;
+ fprintf fmt " rewrite Zpos_xO.\n";
+ fprintf fmt " assert (0 <= Zpos (znz_digits w%i_op)); auto with zarith.\n" size;
+ fprintf fmt " apply F5; auto with arith.\n";
+ fprintf fmt " intros x; case x; clear x; unfold shiftr, same_level.\n";
+ for i = 0 to size do
+ fprintf fmt " intros x y; case y; clear y.\n";
+ for j = 0 to i - 1 do
+ fprintf fmt " intros y; unfold shiftr%i, Ndigits.\n" i;
+ fprintf fmt " repeat rewrite spec_reduce_%i; repeat rewrite spec_reduce_%i; unfold to_Z; intros H1.\n" i j;
+ fprintf fmt " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith.\n" i j i;
+ fprintf fmt " rewrite (spec_zdigits w%i_spec).\n" i;
+ fprintf fmt " rewrite (spec_zdigits w%i_spec).\n" j;
+ fprintf fmt " change (znz_digits w%i_op) with %s.\n" i (genxO (i - j) (" (znz_digits w"^(string_of_int j)^"_op)"));
+ fprintf fmt " repeat rewrite (fun x => Zpos_xO (xO x)).\n";
+ fprintf fmt " repeat rewrite (fun x y => Zpos_xO (%sznz_digits x y)).\n" "@";
+ fprintf fmt " assert (0 <= Zpos (znz_digits w%i_op)); auto with zarith.\n" j;
+ fprintf fmt " try (apply sym_equal; exact (spec_extend%in%i y)).\n" j i;
+
+ done;
+ fprintf fmt " intros y; unfold shiftr%i, Ndigits.\n" i;
+ fprintf fmt " repeat rewrite spec_reduce_%i; unfold to_Z; intros H1.\n" i;
+ fprintf fmt " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith.\n" i i i;
+ for j = i + 1 to size do
+ fprintf fmt " intros y; unfold shiftr%i, Ndigits.\n" j;
+ fprintf fmt " repeat rewrite spec_reduce_%i; repeat rewrite spec_reduce_%i; unfold to_Z; intros H1.\n" i j;
+ fprintf fmt " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith.\n" j j i;
+ fprintf fmt " try (apply sym_equal; exact (spec_extend%in%i x)).\n" i j;
+ done;
+ if i == size then
+ begin
+ fprintf fmt " intros m y; unfold shiftrn, Ndigits.\n";
+ fprintf fmt " repeat rewrite spec_reduce_n; unfold to_Z; intros H1.\n";
+ fprintf fmt " apply F4 with (3:=(wn_spec m))(4:=wn_spec m)(5:=w%i_spec); auto with zarith.\n" size;
+ fprintf fmt " try (apply sym_equal; exact (spec_extend%in m x)).\n" size;
+
+ end
+ else
+ begin
+ fprintf fmt " intros m y; unfold shiftrn, Ndigits.\n";
+ fprintf fmt " repeat rewrite spec_reduce_n; unfold to_Z; intros H1.\n";
+ fprintf fmt " apply F4 with (3:=(wn_spec m))(4:=wn_spec m)(5:=w%i_spec); auto with zarith.\n" i;
+ fprintf fmt " change ([Nn m (extend%i m (extend%i %i x))] = [N%i x]).\n" size i (size - i - 1) i;
+ fprintf fmt " rewrite <- (spec_extend%in m); rewrite <- spec_extend%in%i; auto.\n" size i size;
+ end
+ done;
+ fprintf fmt " intros n x y; case y; clear y;\n";
+ fprintf fmt " intros y; unfold shiftrn, Ndigits; try rewrite spec_reduce_n.\n";
+ for i = 0 to size do
+ fprintf fmt " try rewrite spec_reduce_%i; unfold to_Z; intros H1.\n" i;
+ fprintf fmt " apply F4 with (3:=(wn_spec n))(4:=w%i_spec)(5:=wn_spec n); auto with zarith.\n" i;
+ fprintf fmt " rewrite (spec_zdigits w%i_spec).\n" i;
+ fprintf fmt " rewrite (spec_zdigits (wn_spec n)).\n";
+ fprintf fmt " apply Zle_trans with (2 := F6 n).\n";
+ fprintf fmt " change (znz_digits w%i_op) with %s.\n" size (genxO (size - i) ("(znz_digits w" ^ (string_of_int i) ^ "_op)"));
+ fprintf fmt " repeat rewrite (fun x => Zpos_xO (xO x)).\n";
+ fprintf fmt " repeat rewrite (fun x y => Zpos_xO (%sznz_digits x y)).\n" "@";
+ fprintf fmt " assert (H: 0 <= Zpos (znz_digits w%i_op)); auto with zarith.\n" i;
+ if i == size then
+ fprintf fmt " change ([Nn n (extend%i n y)] = [N%i y]).\n" size i
+ else
+ fprintf fmt " change ([Nn n (extend%i n (extend%i %i y))] = [N%i y]).\n" size i (size - i - 1) i;
+ fprintf fmt " rewrite <- (spec_extend%in n); auto.\n" size;
+ if i <> size then
+ fprintf fmt " try (rewrite <- spec_extend%in%i; auto).\n" i size;
+ done;
+ fprintf fmt " generalize y; clear y; intros m y.\n";
+ fprintf fmt " rewrite spec_reduce_n; unfold to_Z; intros H1.\n";
+ fprintf fmt " apply F4 with (3:=(wn_spec (Max.max n m)))(4:=wn_spec m)(5:=wn_spec n); auto with zarith.\n";
+ fprintf fmt " rewrite (spec_zdigits (wn_spec m)).\n";
+ fprintf fmt " rewrite (spec_zdigits (wn_spec (Max.max n m))).\n";
+ fprintf fmt " apply F5; auto with arith.\n";
+ fprintf fmt " exact (spec_cast_r n m y).\n";
+ fprintf fmt " exact (spec_cast_l n m x).\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+
fprintf fmt " Definition safe_shiftr n x := \n";
fprintf fmt " match compare n (Ndigits x) with\n ";
fprintf fmt " | Lt => shiftr n x \n";
@@ -928,14 +2849,39 @@ let print_Make () =
fprintf fmt " end.\n";
fprintf fmt "\n";
+
+ fprintf fmt " Theorem spec_safe_shiftr: forall n x,\n";
+ fprintf fmt " [safe_shiftr n x] = [x] / 2 ^ [n].\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " intros n x; unfold safe_shiftr;\n";
+ fprintf fmt " generalize (spec_compare n (Ndigits x)); case compare; intros H.\n";
+ fprintf fmt " apply trans_equal with (1 := spec_0 w0_spec).\n";
+ fprintf fmt " apply sym_equal; apply ZDivModAux.Zdiv_lt_0; rewrite H.\n";
+ fprintf fmt " rewrite spec_Ndigits; exact (spec_digits x).\n";
+ fprintf fmt " rewrite <- spec_shiftr; auto with zarith.\n";
+ fprintf fmt " apply trans_equal with (1 := spec_0 w0_spec).\n";
+ fprintf fmt " apply sym_equal; apply ZDivModAux.Zdiv_lt_0.\n";
+ fprintf fmt " rewrite spec_Ndigits in H; case (spec_digits x); intros H1 H2.\n";
+ fprintf fmt " split; auto.\n";
+ fprintf fmt " apply Zlt_le_trans with (1 := H2).\n";
+ fprintf fmt " apply Zpower_le_monotone; auto with zarith.\n";
+ fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+
+ fprintf fmt "\n";
+
(* Shiftl *)
for i = 0 to size do
fprintf fmt " Definition shiftl%i n x := w%i_op.(znz_add_mul_div) n x w%i_op.(znz_0).\n" i i i
done;
fprintf fmt " Definition shiftln n p x := (make_op n).(znz_add_mul_div) p x (make_op n).(znz_0).\n";
- fprintf fmt " Definition shiftl := \n";
- fprintf fmt " Eval lazy beta delta [level] in \n";
- fprintf fmt " level (fun n x => %s0 (shiftl0 n x))\n" c;
+ fprintf fmt " Definition shiftl := Eval lazy beta delta [same_level] in\n";
+ fprintf fmt " same_level _ (fun n x => %s0 (shiftl0 n x))\n" c;
for i = 1 to size do
fprintf fmt " (fun n x => reduce_%i (shiftl%i n x))\n" i i;
done;
@@ -943,1287 +2889,469 @@ let print_Make () =
fprintf fmt "\n";
fprintf fmt "\n";
- (* Double size *)
- fprintf fmt " Definition double_size w := match w with\n";
- fprintf fmt " | %s0 w=> N1 (WW w_0 w)\n" c;
- for i = 1 to size-1 do
- fprintf fmt " | %s%i w=> %s%i (extend1 _ w)\n" c i c (i + 1);
- done;
- fprintf fmt " | %s%i w=> %sn 0 (extend1 _ w)\n" c size c ;
- fprintf fmt " | %sn n w=> %sn (S n) (extend1 _ w)\n" c c;
- fprintf fmt " end.\n";
- fprintf fmt "\n";
-
- (* Safe shiftl *)
- fprintf fmt " Definition safe_shiftl_aux_body cont n x :=\n";
- fprintf fmt " match compare n (head0 x) with\n";
- fprintf fmt " Gt => cont n (double_size x)\n";
- fprintf fmt " | _ => shiftl n x\n";
- fprintf fmt " end.\n";
- fprintf fmt "\n";
- fprintf fmt " Fixpoint safe_shiftl_aux p cont n x {struct p} :=\n";
- fprintf fmt " safe_shiftl_aux_body \n";
- fprintf fmt " (fun n x => match p with\n";
- fprintf fmt " | xH => cont n x\n";
- fprintf fmt " | xO p => safe_shiftl_aux p (safe_shiftl_aux p cont) n x\n";
- fprintf fmt " | xI p => safe_shiftl_aux p (safe_shiftl_aux p cont) n x\n";
- fprintf fmt " end) n x.\n";
- fprintf fmt "\n";
- fprintf fmt " Definition safe_shiftl n x :=\n";
- fprintf fmt " safe_shiftl_aux (digits n) (fun n x => %s0 w0_op.(znz_0)) n x.\n" c;
- fprintf fmt " \n";
-
- (* even *)
- fprintf fmt " Definition is_even x :=\n";
- fprintf fmt " match x with\n";
+ fprintf fmt " Theorem spec_shiftl: forall n x,\n";
+ fprintf fmt " [n] <= [head0 x] -> [shiftl n x] = [x] * 2 ^ [n].\n";
+ if gen_proof then
+ begin
+ fprintf fmt " Proof.\n";
+ fprintf fmt " assert (F0: forall x y, x - (x - y) = y).\n";
+ fprintf fmt " intros; ring.\n";
+ fprintf fmt " assert (F2: forall x y z, 0 <= x -> 0 <= y -> x < z -> 0 <= x / 2 ^ y < z).\n";
+ fprintf fmt " intros x y z HH HH1 HH2.\n";
+ fprintf fmt " split; auto with zarith.\n";
+ fprintf fmt " apply Zle_lt_trans with (2 := HH2); auto with zarith.\n";
+ fprintf fmt " apply ZDivModAux.Zdiv_le_upper_bound; auto with zarith.\n";
+ fprintf fmt " pattern x at 1; replace x with (x * 2 ^ 0); auto with zarith.\n";
+ fprintf fmt " apply Zmult_le_compat_l; auto.\n";
+ fprintf fmt " apply Zpower_le_monotone; auto with zarith.\n";
+ fprintf fmt " rewrite Zpower_exp_0; ring.\n";
+ fprintf fmt " assert (F3: forall x y, 0 <= y -> y <= x -> 0 <= x - y < 2 ^ x).\n";
+ fprintf fmt " intros xx y HH HH1.\n";
+ fprintf fmt " split; auto with zarith.\n";
+ fprintf fmt " apply Zle_lt_trans with xx; auto with zarith.\n";
+ fprintf fmt " apply ZPowerAux.Zpower2_lt_lin; auto with zarith.\n";
+ fprintf fmt " assert (F4: forall ww ww1 ww2 \n";
+ fprintf fmt " (ww_op: znz_op ww) (ww1_op: znz_op ww1) (ww2_op: znz_op ww2)\n";
+ fprintf fmt " xx yy xx1 yy1,\n";
+ fprintf fmt " znz_to_Z ww2_op yy <= znz_to_Z ww1_op (znz_head0 ww1_op xx) ->\n";
+ fprintf fmt " znz_to_Z ww1_op (znz_zdigits ww1_op) <= znz_to_Z ww_op (znz_zdigits ww_op) ->\n";
+ fprintf fmt " znz_spec ww_op -> znz_spec ww1_op -> znz_spec ww2_op ->\n";
+ fprintf fmt " znz_to_Z ww_op xx1 = znz_to_Z ww1_op xx ->\n";
+ fprintf fmt " znz_to_Z ww_op yy1 = znz_to_Z ww2_op yy ->\n";
+ fprintf fmt " znz_to_Z ww_op\n";
+ fprintf fmt " (znz_add_mul_div ww_op yy1\n";
+ fprintf fmt " xx1 (znz_0 ww_op)) = znz_to_Z ww1_op xx * 2 ^ znz_to_Z ww2_op yy).\n";
+ fprintf fmt " intros ww ww1 ww2 ww_op ww1_op ww2_op xx yy xx1 yy1 Hl Hl1 Hw Hw1 Hw2 Hx Hy.\n";
+ fprintf fmt " case (spec_to_Z Hw xx1); auto with zarith; intros HH1 HH2.\n";
+ fprintf fmt " case (spec_to_Z Hw yy1); auto with zarith; intros HH3 HH4.\n";
+ fprintf fmt " rewrite <- Hx.\n";
+ fprintf fmt " rewrite <- Hy.\n";
+ fprintf fmt " generalize (spec_add_mul_div Hw xx1 (znz_0 ww_op) yy1).\n";
+ fprintf fmt " rewrite (spec_0 Hw).\n";
+ fprintf fmt " assert (F1: znz_to_Z ww1_op (znz_head0 ww1_op xx) <= Zpos (znz_digits ww1_op)).\n";
+ fprintf fmt " case (Zle_lt_or_eq _ _ HH1); intros HH5.\n";
+ fprintf fmt " apply Zlt_le_weak.\n";
+ fprintf fmt " case (ZnZ.spec_head0 Hw1 xx).\n";
+ fprintf fmt " rewrite <- Hx; auto.\n";
+ fprintf fmt " intros _ Hu; unfold base in Hu.\n";
+ fprintf fmt " case (Zle_or_lt (Zpos (znz_digits ww1_op))\n";
+ fprintf fmt " (znz_to_Z ww1_op (znz_head0 ww1_op xx))); auto; intros H1.\n";
+ fprintf fmt " absurd (2 ^ (Zpos (znz_digits ww1_op)) <= 2 ^ (znz_to_Z ww1_op (znz_head0 ww1_op xx))).\n";
+ fprintf fmt " apply Zlt_not_le.\n";
+ fprintf fmt " case (spec_to_Z Hw1 xx); intros HHx3 HHx4.\n";
+ fprintf fmt " rewrite <- (Zmult_1_r (2 ^ znz_to_Z ww1_op (znz_head0 ww1_op xx))).\n";
+ fprintf fmt " apply Zle_lt_trans with (2 := Hu).\n";
+ fprintf fmt " apply Zmult_le_compat_l; auto with zarith.\n";
+ fprintf fmt " apply Zpower_le_monotone; auto with zarith.\n";
+ fprintf fmt " rewrite (ZnZ.spec_head00 Hw1 xx); auto with zarith.\n";
+ fprintf fmt " rewrite ZDivModAux.Zdiv_0; auto with zarith.\n";
+ fprintf fmt " rewrite Zplus_0_r.\n";
+ fprintf fmt " case (Zle_lt_or_eq _ _ HH1); intros HH5.\n";
+ fprintf fmt " rewrite Zmod_def_small; auto with zarith.\n";
+ fprintf fmt " intros HH; apply HH.\n";
+ fprintf fmt " rewrite Hy; apply Zle_trans with (1:= Hl).\n";
+ fprintf fmt " rewrite <- (spec_zdigits Hw). \n";
+ fprintf fmt " apply Zle_trans with (2 := Hl1); auto.\n";
+ fprintf fmt " rewrite (spec_zdigits Hw1); auto with zarith.\n";
+ fprintf fmt " split; auto with zarith .\n";
+ fprintf fmt " apply Zlt_le_trans with (base (znz_digits ww1_op)).\n";
+ fprintf fmt " rewrite Hx.\n";
+ fprintf fmt " case (ZnZ.spec_head0 Hw1 xx); auto.\n";
+ fprintf fmt " rewrite <- Hx; auto.\n";
+ fprintf fmt " intros _ Hu; rewrite Zmult_comm in Hu.\n";
+ fprintf fmt " apply Zle_lt_trans with (2 := Hu).\n";
+ fprintf fmt " apply Zmult_le_compat_l; auto with zarith.\n";
+ fprintf fmt " apply Zpower_le_monotone; auto with zarith.\n";
+ fprintf fmt " unfold base; apply Zpower_le_monotone; auto with zarith.\n";
+ fprintf fmt " split; auto with zarith.\n";
+ fprintf fmt " rewrite <- (spec_zdigits Hw); auto with zarith.\n";
+ fprintf fmt " rewrite <- (spec_zdigits Hw1); auto with zarith.\n";
+ fprintf fmt " rewrite <- HH5.\n";
+ fprintf fmt " rewrite Zmult_0_l.\n";
+ fprintf fmt " rewrite Zmod_def_small; auto with zarith.\n";
+ fprintf fmt " intros HH; apply HH.\n";
+ fprintf fmt " rewrite Hy; apply Zle_trans with (1 := Hl).\n";
+ fprintf fmt " rewrite (ZnZ.spec_head00 Hw1 xx); auto with zarith.\n";
+ fprintf fmt " rewrite <- (spec_zdigits Hw); auto with zarith.\n";
+ fprintf fmt " rewrite <- (spec_zdigits Hw1); auto with zarith.\n";
+ fprintf fmt " apply Zpower_lt_0; auto with zarith.\n";
+ fprintf fmt " assert (znz_to_Z ww_op yy1 <= Zpos (znz_digits ww_op)); auto with zarith.\n";
+ fprintf fmt " rewrite Hy; apply Zle_trans with (1 := Hl).\n";
+ fprintf fmt " apply Zle_trans with (1 := F1).\n";
+ fprintf fmt " rewrite <- (spec_zdigits Hw); auto with zarith.\n";
+ fprintf fmt " rewrite <- (spec_zdigits Hw1); auto with zarith.\n";
+ fprintf fmt " assert (F5: forall n m, (n <= m)%snat ->\n" "%";
+ fprintf fmt " Zpos (znz_digits (make_op n)) <= Zpos (znz_digits (make_op m))).\n";
+ fprintf fmt " intros n m HH; elim HH; clear m HH; auto with zarith.\n";
+ fprintf fmt " intros m HH Hrec; apply Zle_trans with (1 := Hrec).\n";
+ fprintf fmt " rewrite make_op_S.\n";
+ fprintf fmt " match goal with |- Zpos ?Y <= ?X => change X with (Zpos (xO Y)) end.\n";
+ fprintf fmt " rewrite Zpos_xO.\n";
+ fprintf fmt " assert (0 <= Zpos (znz_digits (make_op n))); auto with zarith.\n";
+ fprintf fmt " assert (F6: forall n, Zpos (znz_digits w%i_op) <= Zpos (znz_digits (make_op n))).\n" size;
+ fprintf fmt " intros n ; apply Zle_trans with (Zpos (znz_digits (make_op 0))).\n";
+ fprintf fmt " change (znz_digits (make_op 0)) with (xO (znz_digits w%i_op)).\n" size;
+ fprintf fmt " rewrite Zpos_xO.\n";
+ fprintf fmt " assert (0 <= Zpos (znz_digits w%i_op)); auto with zarith.\n" size;
+ fprintf fmt " apply F5; auto with arith.\n";
+ fprintf fmt " intros x; case x; clear x; unfold shiftl, same_level.\n";
for i = 0 to size do
- fprintf fmt " | %s%i wx => w%i_op.(znz_is_even) wx\n" c i i
- done;
- fprintf fmt " | %sn n wx => (make_op n).(znz_is_even) wx\n" c;
- fprintf fmt " end.\n";
- fprintf fmt "\n";
-
-
- fprintf fmt "(* Proof section *)\n";
- fprintf fmt "\n";
+ fprintf fmt " intros x y; case y; clear y.\n";
+ for j = 0 to i - 1 do
+ fprintf fmt " intros y; unfold shiftl%i, head0.\n" i;
+ fprintf fmt " repeat rewrite spec_reduce_%i; repeat rewrite spec_reduce_%i; unfold to_Z; intros H1.\n" i j;
+ fprintf fmt " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith.\n" i j i;
+ fprintf fmt " rewrite (spec_zdigits w%i_spec).\n" i;
+ fprintf fmt " rewrite (spec_zdigits w%i_spec).\n" j;
+ fprintf fmt " change (znz_digits w%i_op) with %s.\n" i (genxO (i - j) (" (znz_digits w"^(string_of_int j)^"_op)"));
+ fprintf fmt " repeat rewrite (fun x => Zpos_xO (xO x)).\n";
+ fprintf fmt " repeat rewrite (fun x y => Zpos_xO (%sznz_digits x y)).\n" "@";
+ fprintf fmt " assert (0 <= Zpos (znz_digits w%i_op)); auto with zarith.\n" j;
+ fprintf fmt " try (apply sym_equal; exact (spec_extend%in%i y)).\n" j i;
+
+ done;
+ fprintf fmt " intros y; unfold shiftl%i, head0.\n" i;
+ fprintf fmt " repeat rewrite spec_reduce_%i; unfold to_Z; intros H1.\n" i;
+ fprintf fmt " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith.\n" i i i;
+ for j = i + 1 to size do
+ fprintf fmt " intros y; unfold shiftl%i, head0.\n" j;
+ fprintf fmt " repeat rewrite spec_reduce_%i; repeat rewrite spec_reduce_%i; unfold to_Z; intros H1.\n" i j;
+ fprintf fmt " apply F4 with (3:=w%i_spec)(4:=w%i_spec)(5:=w%i_spec); auto with zarith.\n" j j i;
+ fprintf fmt " try (apply sym_equal; exact (spec_extend%in%i x)).\n" i j;
+ done;
+ if i == size then
+ begin
+ fprintf fmt " intros m y; unfold shiftln, head0.\n";
+ fprintf fmt " repeat rewrite spec_reduce_n; unfold to_Z; intros H1.\n";
+ fprintf fmt " apply F4 with (3:=(wn_spec m))(4:=wn_spec m)(5:=w%i_spec); auto with zarith.\n" size;
+ fprintf fmt " try (apply sym_equal; exact (spec_extend%in m x)).\n" size;
- if gen_proof1 then
- begin
- fprintf fmt " Let w0_spec: znz_spec w0_op := W0.w_spec.\n";
- for i = 1 to 3 do
- fprintf fmt " Let w%i_spec: znz_spec w%i_op := mk_znz2_spec w%i_spec.\n" i i (i-1)
- done;
- for i = 4 to size + 3 do
- fprintf fmt " Let w%i_spec : znz_spec w%i_op := mk_znz2_karatsuba_spec w%i_spec.\n" i i (i-1)
+ end
+ else
+ begin
+ fprintf fmt " intros m y; unfold shiftln, head0.\n";
+ fprintf fmt " repeat rewrite spec_reduce_n; unfold to_Z; intros H1.\n";
+ fprintf fmt " apply F4 with (3:=(wn_spec m))(4:=wn_spec m)(5:=w%i_spec); auto with zarith.\n" i;
+ fprintf fmt " change ([Nn m (extend%i m (extend%i %i x))] = [N%i x]).\n" size i (size - i - 1) i;
+ fprintf fmt " rewrite <- (spec_extend%in m); rewrite <- spec_extend%in%i; auto.\n" size i size;
+ end
done;
- fprintf fmt "\n";
-
- fprintf fmt " Theorem make_op_S: forall n,\n";
- fprintf fmt " make_op (S n) = mk_zn2z_op_karatsuba (make_op n).\n";
- fprintf fmt " intro n; pattern n; apply lt_wf_ind; clear n.\n";
- fprintf fmt " intros n; case n; clear n.\n";
- fprintf fmt " intros _; unfold make_op, make_op_aux, w%i_op; apply refl_equal.\n" (size + 2);
- fprintf fmt " intros n; case n; clear n.\n";
- fprintf fmt " intros _; unfold make_op, make_op_aux, w%i_op; apply refl_equal.\n" (size + 3);
- fprintf fmt " intros n; case n; clear n.\n";
- fprintf fmt " intros _; unfold make_op, make_op_aux, w%i_op, w%i_op; apply refl_equal.\n" (size + 3) (size + 2);
- fprintf fmt " intros n Hrec.\n";
- fprintf fmt " change (make_op (S (S (S (S n))))) with\n";
- fprintf fmt " (mk_zn2z_op_karatsuba (mk_zn2z_op_karatsuba (mk_zn2z_op_karatsuba (make_op (S n))))).\n";
- fprintf fmt " change (make_op (S (S (S n)))) with\n";
- fprintf fmt " (mk_zn2z_op_karatsuba (mk_zn2z_op_karatsuba (mk_zn2z_op_karatsuba (make_op n)))).\n";
- fprintf fmt " rewrite Hrec; auto with arith.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt " \n";
-
- fprintf fmt " Let wn_spec: forall n, znz_spec (make_op n).\n";
- fprintf fmt " intros n; elim n; clear n.\n";
- fprintf fmt " exact w%i_spec.\n" (size + 1);
- fprintf fmt " intros n Hrec; rewrite make_op_S.\n";
- fprintf fmt " exact (mk_znz2_karatsuba_spec Hrec).\n";
+ fprintf fmt " intros n x y; case y; clear y;\n";
+ fprintf fmt " intros y; unfold shiftln, head0; try rewrite spec_reduce_n.\n";
+ for i = 0 to size do
+ fprintf fmt " try rewrite spec_reduce_%i; unfold to_Z; intros H1.\n" i;
+ fprintf fmt " apply F4 with (3:=(wn_spec n))(4:=w%i_spec)(5:=wn_spec n); auto with zarith.\n" i;
+ fprintf fmt " rewrite (spec_zdigits w%i_spec).\n" i;
+ fprintf fmt " rewrite (spec_zdigits (wn_spec n)).\n";
+ fprintf fmt " apply Zle_trans with (2 := F6 n).\n";
+ fprintf fmt " change (znz_digits w%i_op) with %s.\n" size (genxO (size - i) ("(znz_digits w" ^ (string_of_int i) ^ "_op)"));
+ fprintf fmt " repeat rewrite (fun x => Zpos_xO (xO x)).\n";
+ fprintf fmt " repeat rewrite (fun x y => Zpos_xO (%sznz_digits x y)).\n" "@";
+ fprintf fmt " assert (H: 0 <= Zpos (znz_digits w%i_op)); auto with zarith.\n" i;
+ if i == size then
+ fprintf fmt " change ([Nn n (extend%i n y)] = [N%i y]).\n" size i
+ else
+ fprintf fmt " change ([Nn n (extend%i n (extend%i %i y))] = [N%i y]).\n" size i (size - i - 1) i;
+ fprintf fmt " rewrite <- (spec_extend%in n); auto.\n" size;
+ if i <> size then
+ fprintf fmt " try (rewrite <- spec_extend%in%i; auto).\n" i size;
+ done;
+ fprintf fmt " generalize y; clear y; intros m y.\n";
+ fprintf fmt " repeat rewrite spec_reduce_n; unfold to_Z; intros H1.\n";
+ fprintf fmt " apply F4 with (3:=(wn_spec (Max.max n m)))(4:=wn_spec m)(5:=wn_spec n); auto with zarith.\n";
+ fprintf fmt " rewrite (spec_zdigits (wn_spec m)).\n";
+ fprintf fmt " rewrite (spec_zdigits (wn_spec (Max.max n m))).\n";
+ fprintf fmt " apply F5; auto with arith.\n";
+ fprintf fmt " exact (spec_cast_r n m y).\n";
+ fprintf fmt " exact (spec_cast_l n m x).\n";
fprintf fmt " Qed.\n";
- fprintf fmt " \n";
- end;
-
- fprintf fmt " Open Scope Z_scope.\n";
- fprintf fmt " Notation \"[ x ]\" := (to_Z x).\n";
- fprintf fmt " \n";
-
- if gen_proof2 then
- begin
- for i = 1 to size + 1 do
- fprintf fmt " Let znz_to_Z_%i: forall x y,\n" i;
- fprintf fmt " znz_to_Z w%i_op (WW x y) = \n" i;
- fprintf fmt " znz_to_Z w%i_op x * base (znz_digits w%i_op) + znz_to_Z w%i_op y.\n" (i-1) (i-1) (i-1);
- fprintf fmt " Proof.\n";
- fprintf fmt " auto.\n";
- fprintf fmt " Qed. \n";
- fprintf fmt "\n";
- done;
-
- fprintf fmt " Let znz_to_Z_n: forall n x y,\n";
- fprintf fmt " znz_to_Z (make_op (S n)) (WW x y) = \n";
- fprintf fmt " znz_to_Z (make_op n) x * base (znz_digits (make_op n)) + znz_to_Z (make_op n) y.\n";
- fprintf fmt " Proof.\n";
- fprintf fmt " intros n x y; rewrite make_op_S; auto.\n";
- fprintf fmt " Qed. \n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
fprintf fmt "\n";
- end;
- fprintf fmt " Theorem succ_spec: forall n, [succ n] = [n] + 1.\n";
- if gen_proof3 then
- begin
- fprintf fmt " Proof.\n";
- fprintf fmt " intros n; case n; unfold succ, to_Z.\n";
- for i = 0 to size do
- fprintf fmt " intros n1; generalize (spec_succ_c w%i_spec n1);\n" i;
- fprintf fmt " unfold succ, to_Z, w%i_succ_c; case znz_succ_c; auto.\n" i;
- fprintf fmt " intros ww H; rewrite <- H.\n";
- fprintf fmt " (rewrite znz_to_Z_%i; unfold interp_carry;\n" (i + 1);
- fprintf fmt " apply f_equal2 with (f := Zplus); auto;\n";
- fprintf fmt " apply f_equal2 with (f := Zmult); auto;\n";
- fprintf fmt " exact (spec_1 w%i_spec)).\n" i;
+ (* Double size *)
+ fprintf fmt " Definition double_size w := match w with\n";
+ for i = 0 to size-1 do
+ fprintf fmt " | %s%i x => %s%i (WW (znz_0 w%i_op) x)\n" c i c (i + 1) i;
done;
- fprintf fmt " intros k n1; generalize (spec_succ_c (wn_spec k) n1).\n";
- fprintf fmt " unfold succ, to_Z; case znz_succ_c; auto.\n";
- fprintf fmt " intros ww H; rewrite <- H.\n";
- fprintf fmt " (rewrite (znz_to_Z_n k); unfold interp_carry;\n";
- fprintf fmt " apply f_equal2 with (f := Zplus); auto;\n";
- fprintf fmt " apply f_equal2 with (f := Zmult); auto;\n";
- fprintf fmt " exact (spec_1 (wn_spec k))).\n";
- fprintf fmt " Qed.\n ";
- end else
- fprintf fmt " Admitted.\n";
+ fprintf fmt " | %s%i x => %sn 0 (WW (znz_0 w%i_op) x)\n" c size c size;
+ fprintf fmt " | %sn n x => %sn (S n) (WW (znz_0 (make_op n)) x)\n" c c;
+ fprintf fmt " end.\n";
fprintf fmt "\n";
- if gen_proof4 then
+ fprintf fmt " Theorem spec_double_size_digits: \n";
+ fprintf fmt " forall x, digits (double_size x) = xO (digits x).\n";
+ if gen_proof then
begin
- for i = 0 to size do
- fprintf fmt " Let spec_w%i_add: forall x y, [w%i_add x y] = [%s%i x] + [%s%i y].\n" i i c i c i;
- fprintf fmt " Proof.\n";
- fprintf fmt " intros n m; unfold to_Z, w%i_add, w%i_add_c.\n" i i;
- fprintf fmt " generalize (spec_add_c w%i_spec n m); case znz_add_c; auto.\n" i;
- fprintf fmt " intros ww H; rewrite <- H.\n";
- fprintf fmt " rewrite znz_to_Z_%i; unfold interp_carry;\n" (i + 1);
- fprintf fmt " apply f_equal2 with (f := Zplus); auto;\n";
- fprintf fmt " apply f_equal2 with (f := Zmult); auto;\n";
- fprintf fmt " exact (spec_1 w%i_spec).\n" i;
- fprintf fmt " Qed.\n";
- fprintf fmt " Hint Rewrite spec_w%i_add: addr.\n" i;
- fprintf fmt "\n";
- done;
- fprintf fmt " Let spec_wn_add: forall n x y, [addn n x y] = [%sn n x] + [%sn n y].\n" c c;
fprintf fmt " Proof.\n";
- fprintf fmt " intros k n m; unfold to_Z, addn.\n";
- fprintf fmt " generalize (spec_add_c (wn_spec k) n m); case znz_add_c; auto.\n";
- fprintf fmt " intros ww H; rewrite <- H.\n";
- fprintf fmt " rewrite (znz_to_Z_n k); unfold interp_carry;\n";
- fprintf fmt " apply f_equal2 with (f := Zplus); auto;\n";
- fprintf fmt " apply f_equal2 with (f := Zmult); auto;\n";
- fprintf fmt " exact (spec_1 (wn_spec k)).\n";
+ fprintf fmt " intros x; case x; unfold double_size, digits; clear x; auto.\n";
+ fprintf fmt " intros n x; rewrite make_op_S; auto.\n";
fprintf fmt " Qed.\n";
- fprintf fmt " Hint Rewrite spec_wn_add: addr.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
fprintf fmt "\n";
- for i = 0 to size do
- fprintf fmt " Let spec_w%i_eq0: forall x, if w%i_eq0 x then [%s%i x] = 0 else True.\n" i i c i;
- fprintf fmt " Proof.\n";
- fprintf fmt " intros x; unfold w%i_eq0.\n" i;
- fprintf fmt " generalize (spec_eq0 w%i_spec x); case znz_eq0; auto.\n" i;
- fprintf fmt " Qed.\n";
- fprintf fmt "\n";
- done;
- fprintf fmt " Let spec_extendn_0: forall n wx, [%sn n (extend n _ wx)] = [%sn 0 wx].\n" c c;
- fprintf fmt " intros n; elim n; auto.\n";
- fprintf fmt " intros n1 Hrec wx; simpl extend; rewrite <- Hrec; auto.\n";
- fprintf fmt " unfold to_Z.\n";
- fprintf fmt " case n1; auto; intros n2; repeat rewrite make_op_S; auto.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt " Hint Rewrite spec_extendn_0: extr.\n";
- fprintf fmt "\n";
- fprintf fmt " Let spec_extendn0_0: forall n wx, [%sn (S n) (WW W0 wx)] = [%sn n wx].\n" c c;
- fprintf fmt " Proof.\n";
- fprintf fmt " intros n x; unfold to_Z.\n";
- fprintf fmt " rewrite znz_to_Z_n.\n";
- fprintf fmt " rewrite <- (Zplus_0_l (znz_to_Z (make_op n) x)).\n";
- fprintf fmt " apply (f_equal2 Zplus); auto.\n";
- fprintf fmt " case n; auto.\n";
- fprintf fmt " intros n1; rewrite make_op_S; auto.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt " Hint Rewrite spec_extendn_0: extr.\n";
- fprintf fmt "\n";
- fprintf fmt " Let spec_extend_tr: forall m n (w: word _ (S n)),\n";
- fprintf fmt " [%sn (m + n) (extend_tr w m)] = [%sn n w].\n" c c;
- fprintf fmt " Proof.\n";
- fprintf fmt " induction m; auto.\n";
- fprintf fmt " intros n x; simpl extend_tr.\n";
- fprintf fmt " simpl plus; rewrite spec_extendn0_0; auto.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt " Hint Rewrite spec_extend_tr: extr.\n";
- fprintf fmt "\n";
- fprintf fmt " Let spec_cast_l: forall n m x1,\n";
- fprintf fmt " [%sn (Max.max n m)\n" c;
- fprintf fmt " (castm (diff_r n m) (extend_tr x1 (snd (diff n m))))] =\n";
- fprintf fmt " [%sn n x1].\n" c;
- fprintf fmt " Proof.\n";
- fprintf fmt " intros n m x1; case (diff_r n m); simpl castm.\n";
- fprintf fmt " rewrite spec_extend_tr; auto.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt " Hint Rewrite spec_cast_l: extr.\n";
- fprintf fmt "\n";
- fprintf fmt " Let spec_cast_r: forall n m x1,\n";
- fprintf fmt " [%sn (Max.max n m)\n" c;
- fprintf fmt " (castm (diff_l n m) (extend_tr x1 (fst (diff n m))))] =\n";
- fprintf fmt " [%sn m x1].\n" c;
+ fprintf fmt " Theorem spec_double_size: forall x, [double_size x] = [x].\n";
+ if gen_proof then
+ begin
fprintf fmt " Proof.\n";
- fprintf fmt " intros n m x1; case (diff_l n m); simpl castm.\n";
- fprintf fmt " rewrite spec_extend_tr; auto.\n";
+ fprintf fmt " intros x; case x; unfold double_size; clear x.\n";
+ for i = 0 to size do
+ fprintf fmt " intros x; unfold to_Z, make_op; \n";
+ fprintf fmt " rewrite znz_to_Z_%i; rewrite (spec_0 w%i_spec); auto with zarith.\n" (i + 1) i;
+ done;
+ fprintf fmt " intros n x; unfold to_Z;\n";
+ fprintf fmt " generalize (znz_to_Z_n n); simpl word.\n";
+ fprintf fmt " intros HH; rewrite HH; clear HH.\n";
+ fprintf fmt " generalize (spec_0 (wn_spec n)); simpl word.\n";
+ fprintf fmt " intros HH; rewrite HH; clear HH; auto with zarith.\n";
fprintf fmt " Qed.\n";
- fprintf fmt " Hint Rewrite spec_cast_r: extr.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
fprintf fmt "\n";
- fprintf fmt " Let spec_extend0_0: forall wx, [%s1 (WW w_0 wx)] = [%s0 wx].\n" c c;
- fprintf fmt " Proof.\n";
- fprintf fmt " intros x; unfold to_Z.\n";
- fprintf fmt " rewrite <- (Zplus_0_l (znz_to_Z w0_op x)).\n";
- fprintf fmt " rewrite znz_to_Z_1.\n";
- fprintf fmt " rewrite <- (Zmult_0_l (base (znz_digits w0_op))).\n";
- fprintf fmt " apply (f_equal2 Zplus); auto.\n";
- fprintf fmt " apply (f_equal2 Zmult); auto.\n";
- fprintf fmt " exact (spec_0 w0_spec); auto.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt " Hint Rewrite spec_extend0_0: extr.\n";
- fprintf fmt " \n";
-
- for i = 1 to size do
- for j = 1 to size - i do
- fprintf fmt " Let spec_extend%i_%i: forall wx, [%s%i (extend%i _ wx)] = [%s%i wx].\n" i j c (i + j) i c j;
- fprintf fmt " Proof.
- intros x; unfold extend%i, to_Z.\n" i;
- fprintf fmt " rewrite <- (Zplus_0_l (znz_to_Z w%i_op x)).\n" j;
- fprintf fmt " rewrite znz_to_Z_%i; auto.\n" (i + j);
- fprintf fmt " Qed.\n";
- fprintf fmt " Hint Rewrite spec_extend%i_%i: extr.\n" i j;
- fprintf fmt "\n";
- done;
- fprintf fmt " Let spec_extend%i_0: forall wx, [%sn 0 (extend%i _ wx)] = [N%i wx].\n" i c i (size + 1 - i);
- fprintf fmt " Proof.\n";
- fprintf fmt " intros x; unfold extend%i, to_Z; auto.\n" (size + 1 - i);
- fprintf fmt " Qed.\n";
- fprintf fmt " Hint Rewrite spec_extend%i_0: extr.\n" i;
- fprintf fmt " \n";
- done;
- end;
-
- fprintf fmt " Theorem spec_add: forall x y, [add x y] = [x] + [y].\n";
- if gen_proof5 then
+ fprintf fmt " Theorem spec_double_size_head0: \n";
+ fprintf fmt " forall x, 2 * [head0 x] <= [head0 (double_size x)].\n";
+ if gen_proof then
begin
fprintf fmt " Proof.\n";
- fprintf fmt " intros x y; case x; unfold add.\n";
- fprintf fmt " intros x1; case y.\n";
- fprintf fmt " intros y1; rewrite spec_w0_add; auto.\n";
- for i = 1 to size do
- fprintf fmt " intros y1; generalize (spec_w0_eq0 x1); case w0_eq0; intros HH.\n";
- fprintf fmt " rewrite HH; auto.\n";
- if i = 1 then
- fprintf fmt " rewrite spec_w1_add; rewrite spec_extend0_0; auto.\n"
- else
- fprintf fmt " rewrite spec_w%i_add; rewrite spec_extend%i_1; rewrite spec_extend0_0; auto.\n" i (i -1);
- done;
- fprintf fmt " intros n y1; generalize (spec_w0_eq0 x1); case w0_eq0; intros HH.\n";
- fprintf fmt " rewrite HH; auto.\n";
- fprintf fmt " rewrite spec_wn_add.\n";
- fprintf fmt " rewrite spec_extendn_0; rewrite spec_extend%i_0; rewrite spec_extend0_0; auto.\n" size;
- for i = 1 to size do
- fprintf fmt " intros x1; case y.\n";
- fprintf fmt " intros y1; generalize (spec_w0_eq0 y1); case w0_eq0; intros HH.\n";
- fprintf fmt " rewrite HH; rewrite Zplus_0_r; auto.\n";
- if i = 1 then
- fprintf fmt " rewrite spec_w1_add; rewrite spec_extend0_0; auto.\n"
- else
- fprintf fmt " rewrite spec_w%i_add; rewrite spec_extend%i_1; rewrite spec_extend0_0; auto.\n" i (i-1);
- for j = 1 to size do
- if i <= j then
- fprintf fmt " intros y1; rewrite spec_w%i_add; auto.\n" j
- else
- fprintf fmt " intros y1; rewrite spec_w%i_add; auto.\n" i;
- done;
- fprintf fmt " intros n y1; rewrite spec_wn_add.\n";
- fprintf fmt " rewrite spec_extendn_0; rewrite spec_extend%i_0; auto.\n" (size + 1 - i);
- done;
- fprintf fmt " intros n x1; case y.\n";
- fprintf fmt " intros y1; generalize (spec_w0_eq0 y1); case w0_eq0; intros HH.\n";
- fprintf fmt " rewrite HH; rewrite Zplus_0_r; auto.\n";
- fprintf fmt " rewrite spec_wn_add; rewrite spec_extendn_0; \n";
- fprintf fmt " rewrite spec_extend%i_0; rewrite spec_extend0_0; auto.\n" size;
- for i = 1 to size do
- fprintf fmt " intros y1; rewrite spec_wn_add; rewrite spec_extendn_0; rewrite spec_extend%i_0; auto.\n" (size + 1 - i);
- done;
- fprintf fmt " intros m y1; rewrite spec_wn_add; rewrite spec_cast_l; rewrite spec_cast_r; auto.\n";
+ fprintf fmt " intros x.\n";
+ fprintf fmt " assert (F1:= to_Z_pos (head0 x)).\n";
+ fprintf fmt " assert (F2: 0 < Zpos (digits x)).\n";
+ fprintf fmt " red; auto.\n";
+ fprintf fmt " case (Zle_lt_or_eq _ _ (to_Z_pos x)); intros HH.\n";
+ fprintf fmt " generalize HH; rewrite <- (spec_double_size x); intros HH1.\n";
+ fprintf fmt " case (spec_head0 x HH); intros _ HH2.\n";
+ fprintf fmt " case (spec_head0 _ HH1).\n";
+ fprintf fmt " rewrite (spec_double_size x); rewrite (spec_double_size_digits x).\n";
+ fprintf fmt " intros HH3 _.\n";
+ fprintf fmt " case (Zle_or_lt ([head0 (double_size x)]) (2 * [head0 x])); auto; intros HH4.\n";
+ fprintf fmt " absurd (2 ^ (2 * [head0 x] )* [x] < 2 ^ [head0 (double_size x)] * [x]); auto.\n";
+ fprintf fmt " apply Zle_not_lt.\n";
+ fprintf fmt " apply Zmult_le_compat_r; auto with zarith.\n";
+ fprintf fmt " apply Zpower_le_monotone; auto; auto with zarith.\n";
+ fprintf fmt " generalize (to_Z_pos (head0 (double_size x))); auto with zarith.\n";
+ fprintf fmt " assert (HH5: 2 ^[head0 x] <= 2 ^(Zpos (digits x) - 1)).\n";
+ fprintf fmt " case (Zle_lt_or_eq 1 [x]); auto with zarith; intros HH5.\n";
+ fprintf fmt " apply Zmult_le_reg_r with (2 ^ 1); auto with zarith.\n";
+ fprintf fmt " rewrite <- (fun x y z => Zpower_exp x (y - z)); auto with zarith.\n";
+ fprintf fmt " assert (tmp: forall x, x - 1 + 1 = x); [intros; ring | rewrite tmp; clear tmp].\n";
+ fprintf fmt " apply Zle_trans with (2 := Zlt_le_weak _ _ HH2).\n";
+ fprintf fmt " apply Zmult_le_compat_l; auto with zarith.\n";
+ fprintf fmt " rewrite Zpower_exp_1; auto with zarith.\n";
+ fprintf fmt " apply Zpower_le_monotone; auto with zarith.\n";
+ fprintf fmt " split; auto with zarith. \n";
+ fprintf fmt " case (Zle_or_lt (Zpos (digits x)) [head0 x]); auto with zarith; intros HH6.\n";
+ fprintf fmt " absurd (2 ^ Zpos (digits x) <= 2 ^ [head0 x] * [x]); auto with zarith.\n";
+ fprintf fmt " rewrite <- HH5; rewrite Zmult_1_r.\n";
+ fprintf fmt " apply Zpower_le_monotone; auto with zarith.\n";
+ fprintf fmt " rewrite (Zmult_comm 2).\n";
+ fprintf fmt " rewrite Zpower_mult; auto with zarith.\n";
+ fprintf fmt " rewrite Zpower_2.\n";
+ fprintf fmt " apply Zlt_le_trans with (2 := HH3).\n";
+ fprintf fmt " rewrite <- Zmult_assoc.\n";
+ fprintf fmt " replace (Zpos (xO (digits x)) - 1) with\n";
+ fprintf fmt " ((Zpos (digits x) - 1) + (Zpos (digits x))).\n";
+ fprintf fmt " rewrite Zpower_exp; auto with zarith.\n";
+ fprintf fmt " apply Zmult_lt_compat; auto with zarith.\n";
+ fprintf fmt " split; auto with zarith.\n";
+ fprintf fmt " apply Zmult_lt_0_compat; auto with zarith.\n";
+ fprintf fmt " rewrite Zpos_xO; ring.\n";
+ fprintf fmt " apply Zlt_le_weak; auto.\n";
+ fprintf fmt " repeat rewrite spec_head00; auto.\n";
+ fprintf fmt " rewrite spec_double_size_digits.\n";
+ fprintf fmt " rewrite Zpos_xO; auto with zarith.\n";
+ fprintf fmt " rewrite spec_double_size; auto.\n";
fprintf fmt " Qed.\n";
- end else
+ end
+ else
fprintf fmt " Admitted.\n";
fprintf fmt "\n";
- if gen_proof6 then
+ fprintf fmt " Theorem spec_double_size_head0_pos: \n";
+ fprintf fmt " forall x, 0 < [head0 (double_size x)].\n";
+ if gen_proof then
begin
- fprintf fmt " Let spec_reduce_0: forall x, [reduce_0 x] = [%s0 x].\n" c;
fprintf fmt " Proof.\n";
- fprintf fmt " intros x; unfold to_Z, reduce_0.\n";
- fprintf fmt " auto.\n";
+ fprintf fmt " intros x.\n";
+ fprintf fmt " assert (F: 0 < Zpos (digits x)).\n";
+ fprintf fmt " red; auto.\n";
+ fprintf fmt " case (Zle_lt_or_eq _ _ (to_Z_pos (head0 (double_size x)))); auto; intros F0.\n";
+ fprintf fmt " case (Zle_lt_or_eq _ _ (to_Z_pos (head0 x))); intros F1.\n";
+ fprintf fmt " apply Zlt_le_trans with (2 := (spec_double_size_head0 x)); auto with zarith.\n";
+ fprintf fmt " case (Zle_lt_or_eq _ _ (to_Z_pos x)); intros F3.\n";
+ fprintf fmt " generalize F3; rewrite <- (spec_double_size x); intros F4.\n";
+ fprintf fmt " absurd (2 ^ (Zpos (xO (digits x)) - 1) < 2 ^ (Zpos (digits x))).\n";
+ fprintf fmt " apply Zle_not_lt.\n";
+ fprintf fmt " apply Zpower_le_monotone; auto with zarith.\n";
+ fprintf fmt " split; auto with zarith.\n";
+ fprintf fmt " rewrite Zpos_xO; auto with zarith.\n";
+ fprintf fmt " case (spec_head0 x F3).\n";
+ fprintf fmt " rewrite <- F1; rewrite Zpower_exp_0; rewrite Zmult_1_l; intros _ HH.\n";
+ fprintf fmt " apply Zle_lt_trans with (2 := HH).\n";
+ fprintf fmt " case (spec_head0 _ F4).\n";
+ fprintf fmt " rewrite (spec_double_size x); rewrite (spec_double_size_digits x).\n";
+ fprintf fmt " rewrite <- F0; rewrite Zpower_exp_0; rewrite Zmult_1_l; auto.\n";
+ fprintf fmt " generalize F1; rewrite (spec_head00 _ (sym_equal F3)); auto with zarith.\n";
fprintf fmt " Qed.\n";
- fprintf fmt " \n";
-
- for i = 1 to size + 1 do
- if (i == size + 1) then
- fprintf fmt " Let spec_reduce_%i: forall x, [reduce_%i x] = [%sn 0 x].\n" i i c
- else
- fprintf fmt " Let spec_reduce_%i: forall x, [reduce_%i x] = [%s%i x].\n" i i c i;
- fprintf fmt " Proof.\n";
- fprintf fmt " intros x; case x; unfold reduce_%i.\n" i;
- fprintf fmt " exact (spec_0 w0_spec).\n";
- fprintf fmt " intros x1 y1.\n";
- fprintf fmt " generalize (spec_w%i_eq0 x1); \n" (i - 1);
- fprintf fmt " case w%i_eq0; intros H1; auto.\n" (i - 1);
- if i <> 1 then
- fprintf fmt " rewrite spec_reduce_%i.\n" (i - 1);
- fprintf fmt " unfold to_Z; rewrite znz_to_Z_%i.\n" i;
- fprintf fmt " unfold to_Z in H1; rewrite H1; auto.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt " \n";
- done;
-
- fprintf fmt " Let spec_reduce_n: forall n x, [reduce_n n x] = [%sn n x].\n" c;
- fprintf fmt " Proof.\n";
- fprintf fmt " intros n; elim n; simpl reduce_n.\n";
- fprintf fmt " intros x; rewrite <- spec_reduce_%i; auto.\n" (size + 1);
- fprintf fmt " intros n1 Hrec x; case x.\n";
- fprintf fmt " unfold to_Z; rewrite make_op_S; auto.\n";
- fprintf fmt " exact (spec_0 w0_spec).\n";
- fprintf fmt " intros x1 y1; case x1; auto.\n";
- fprintf fmt " rewrite Hrec.\n";
- fprintf fmt " rewrite spec_extendn0_0; auto.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt " \n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
- fprintf fmt " Let to_Z_pos: forall x, 0 <= [x].\n";
- fprintf fmt " Proof.\n";
- fprintf fmt " intros x; case x; unfold to_Z.\n";
- for i = 0 to size do
- fprintf fmt " intros x1; case (spec_to_Z w%i_spec x1); auto.\n" i;
- done;
- fprintf fmt " intros n x1; case (spec_to_Z (wn_spec n) x1); auto.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt " \n";
- fprintf fmt " Let spec_pred: forall x, 0 < [x] -> [pred x] = [x] - 1.\n";
- fprintf fmt " Proof.\n";
- fprintf fmt " intros x; case x; unfold pred.\n";
- for i = 0 to size do
- fprintf fmt " intros x1 H1; unfold w%i_pred_c; \n" i;
- fprintf fmt " generalize (spec_pred_c w%i_spec x1); case znz_pred_c; intros y1.\n" i;
- fprintf fmt " rewrite spec_reduce_%i; auto.\n" i;
- fprintf fmt " unfold interp_carry; unfold to_Z.\n";
- fprintf fmt " case (spec_to_Z w%i_spec x1); intros HH1 HH2.\n" i;
- fprintf fmt " case (spec_to_Z w%i_spec y1); intros HH3 HH4 HH5.\n" i;
- fprintf fmt " assert (znz_to_Z w%i_op x1 - 1 < 0); auto with zarith.\n" i;
- fprintf fmt " unfold to_Z in H1; auto with zarith.\n";
- done;
- fprintf fmt " intros n x1 H1; \n";
- fprintf fmt " generalize (spec_pred_c (wn_spec n) x1); case znz_pred_c; intros y1.\n";
- fprintf fmt " rewrite spec_reduce_n; auto.\n";
- fprintf fmt " unfold interp_carry; unfold to_Z.\n";
- fprintf fmt " case (spec_to_Z (wn_spec n) x1); intros HH1 HH2.\n";
- fprintf fmt " case (spec_to_Z (wn_spec n) y1); intros HH3 HH4 HH5.\n";
- fprintf fmt " assert (znz_to_Z (make_op n) x1 - 1 < 0); auto with zarith.\n";
- fprintf fmt " unfold to_Z in H1; auto with zarith.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt " \n";
-
- fprintf fmt " Let spec_pred0: forall x, [x] = 0 -> [pred x] = 0.\n";
- fprintf fmt " Proof.\n";
- fprintf fmt " intros x; case x; unfold pred.\n";
- for i = 0 to size do
- fprintf fmt " intros x1 H1; unfold w%i_pred_c; \n" i;
- fprintf fmt " generalize (spec_pred_c w%i_spec x1); case znz_pred_c; intros y1.\n" i;
- fprintf fmt " unfold interp_carry; unfold to_Z.\n";
- fprintf fmt " unfold to_Z in H1; auto with zarith.\n";
- fprintf fmt " case (spec_to_Z w%i_spec y1); intros HH3 HH4; auto with zarith.\n" i;
- fprintf fmt " intros; exact (spec_0 w0_spec).\n";
- done;
- fprintf fmt " intros n x1 H1; \n";
- fprintf fmt " generalize (spec_pred_c (wn_spec n) x1); case znz_pred_c; intros y1.\n";
- fprintf fmt " unfold interp_carry; unfold to_Z.\n";
- fprintf fmt " unfold to_Z in H1; auto with zarith.\n";
- fprintf fmt " case (spec_to_Z (wn_spec n) y1); intros HH3 HH4; auto with zarith.\n";
- fprintf fmt " intros; exact (spec_0 w0_spec).\n";
- fprintf fmt " Qed.\n";
- fprintf fmt " \n";
-
- for i = 0 to size do
- fprintf fmt " Let spec_w%i_sub: forall x y, [%s%i y] <= [%s%i x] -> [w%i_sub x y] = [%s%i x] - [%s%i y].\n" i c i c i i c i c i;
- fprintf fmt " Proof.\n";
- fprintf fmt " intros n m; unfold w%i_sub, w%i_sub_c.\n" i i;
- fprintf fmt " generalize (spec_sub_c w%i_spec n m); case znz_sub_c; \n" i;
- if i == 0 then
- fprintf fmt " intros x; auto.\n"
- else
- fprintf fmt " intros x; try rewrite spec_reduce_%i; auto.\n" i;
- fprintf fmt " unfold interp_carry; unfold zero, w_0, to_Z.\n";
- fprintf fmt " rewrite (spec_0 w0_spec).\n";
- fprintf fmt " case (spec_to_Z w%i_spec x); intros; auto with zarith.\n" i;
- fprintf fmt " Qed.\n";
- fprintf fmt "\n";
- done;
+ (* Safe shiftl *)
- fprintf fmt " Let spec_wn_sub: forall n x y, [%sn n y] <= [%sn n x] -> [subn n x y] = [%sn n x] - [%sn n y].\n" c c c c;
- fprintf fmt " Proof.\n";
- fprintf fmt " intros k n m; unfold subn.\n";
- fprintf fmt " generalize (spec_sub_c (wn_spec k) n m); case znz_sub_c; \n";
- fprintf fmt " intros x; auto.\n";
- fprintf fmt " unfold interp_carry, to_Z.\n";
- fprintf fmt " case (spec_to_Z (wn_spec k) x); intros; auto with zarith.\n";
- fprintf fmt " Qed.\n";
+ fprintf fmt " Definition safe_shiftl_aux_body cont n x :=\n";
+ fprintf fmt " match compare n (head0 x) with\n";
+ fprintf fmt " Gt => cont n (double_size x)\n";
+ fprintf fmt " | _ => shiftl n x\n";
+ fprintf fmt " end.\n";
fprintf fmt "\n";
- end;
- fprintf fmt " Theorem spec_sub: forall x y, [y] <= [x] -> [sub x y] = [x] - [y].\n";
- if gen_proof7 then
+ fprintf fmt " Theorem spec_safe_shift_aux_body: forall n p x cont,\n";
+ fprintf fmt " 2^ Zpos p <= [head0 x] ->\n";
+ fprintf fmt " (forall x, 2 ^ (Zpos p + 1) <= [head0 x]->\n";
+ fprintf fmt " [cont n x] = [x] * 2 ^ [n]) ->\n";
+ fprintf fmt " [safe_shiftl_aux_body cont n x] = [x] * 2 ^ [n].\n";
+ if gen_proof then
begin
fprintf fmt " Proof.\n";
- fprintf fmt " intros x y; case x; unfold sub.\n";
- fprintf fmt " intros x1; case y.\n";
- fprintf fmt " intros y1 H; rewrite spec_w0_sub; auto.\n";
- for i = 1 to size do
- fprintf fmt " intros y1 H; generalize (spec_w0_eq0 x1); case w0_eq0; intros HH.\n";
- fprintf fmt " generalize H; rewrite HH; unfold to_Z, zero, w_0.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); case (spec_to_Z w%i_spec y1); auto with zarith.\n" i;
- if i == 1 then
- fprintf fmt " rewrite spec_w1_sub; rewrite spec_extend0_0; auto.\n"
- else
- fprintf fmt " rewrite spec_w%i_sub; rewrite spec_extend%i_1; rewrite spec_extend0_0; auto.\n" i (i - 1);
- done;
- fprintf fmt " intros n y1 H; generalize (spec_w0_eq0 x1); case w0_eq0; intros HH.\n";
- fprintf fmt " generalize H; rewrite HH; unfold to_Z, zero, w_0.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); case (spec_to_Z (wn_spec n) y1); auto with zarith.\n";
- fprintf fmt " rewrite spec_wn_sub; rewrite spec_extendn_0; rewrite spec_extend%i_0; rewrite spec_extend0_0; auto.\n" size;
- for i = 1 to size do
- fprintf fmt " intros x1; case y.\n";
- fprintf fmt " intros y1 H; generalize (spec_w0_eq0 y1); case w0_eq0; intros HH.\n";
- fprintf fmt " rewrite HH; rewrite Zminus_0_r; auto.\n";
- if i = 1 then
- fprintf fmt " rewrite spec_w1_sub; rewrite spec_extend0_0; auto.\n"
- else
- fprintf fmt " rewrite spec_w%i_sub; rewrite spec_extend%i_1; rewrite spec_extend0_0; auto.\n" i (i-1);
- for j = 1 to size do
- if i <= j then
- fprintf fmt " intros y1 H; rewrite spec_w%i_sub; auto.\n" j
- else
- fprintf fmt " intros y1 H; rewrite spec_w%i_sub; auto.\n" i;
- done;
- fprintf fmt " intros n y1 H; rewrite spec_wn_sub;\n";
- fprintf fmt " rewrite spec_extendn_0; rewrite spec_extend%i_0; auto.\n" (size + 1 - i);
- done;
- fprintf fmt " intros n x1; case y.\n";
- fprintf fmt " intros y1 H; generalize (spec_w0_eq0 y1); case w0_eq0; intros HH.\n";
- fprintf fmt " rewrite HH; rewrite Zminus_0_r; auto.\n";
- fprintf fmt " rewrite spec_wn_sub; rewrite spec_extendn_0; \n";
- fprintf fmt " rewrite spec_extend%i_0; rewrite spec_extend0_0; auto.\n" size;
- for i = 1 to size do
- fprintf fmt " intros y1 H; rewrite spec_wn_sub; rewrite spec_extendn_0; rewrite spec_extend%i_0; auto.\n" (size + 1 - i);
- done;
- fprintf fmt " intros m y1 H; rewrite spec_wn_sub; rewrite spec_cast_l; rewrite spec_cast_r; auto.\n";
+ fprintf fmt " intros n p x cont H1 H2; unfold safe_shiftl_aux_body.\n";
+ fprintf fmt " generalize (spec_compare n (head0 x)); case compare; intros H.\n";
+ fprintf fmt " apply spec_shiftl; auto with zarith.\n";
+ fprintf fmt " apply spec_shiftl; auto with zarith.\n";
+ fprintf fmt " rewrite H2.\n";
+ fprintf fmt " rewrite spec_double_size; auto.\n";
+ fprintf fmt " rewrite Zplus_comm; rewrite Zpower_exp; auto with zarith.\n";
+ fprintf fmt " apply Zle_trans with (2 := spec_double_size_head0 x).\n";
+ fprintf fmt " rewrite Zpower_exp_1; apply Zmult_le_compat_l; auto with zarith.\n";
fprintf fmt " Qed.\n";
- end else
+ end
+ else
fprintf fmt " Admitted.\n";
fprintf fmt "\n";
- if gen_proof8 then
- begin
- for i = 0 to size do
- fprintf fmt " Let spec_w%i_sub0: forall x y, [%s%i x] < [%s%i y] -> [w%i_sub x y] = 0.\n" i c i c i i;
- fprintf fmt " Proof.\n";
- fprintf fmt " intros n m; unfold w%i_sub, w%i_sub_c.\n" i i;
- fprintf fmt " generalize (spec_sub_c w%i_spec n m); case znz_sub_c; \n" i;
- fprintf fmt " intros x; unfold interp_carry.\n";
- fprintf fmt " unfold to_Z; case (spec_to_Z w%i_spec x); intros; auto with zarith.\n" i;
- fprintf fmt " intros; unfold to_Z, zero, w_0; rewrite (spec_0 w0_spec); auto.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt "\n";
- done;
-
- fprintf fmt " Let spec_wn_sub0: forall n x y, [%sn n x] < [%sn n y] -> [subn n x y] = 0.\n" c c;
- fprintf fmt " Proof.\n";
- fprintf fmt " intros k n m; unfold subn.\n";
- fprintf fmt " generalize (spec_sub_c (wn_spec k) n m); case znz_sub_c; \n";
- fprintf fmt " intros x; unfold interp_carry.\n";
- fprintf fmt " unfold to_Z; case (spec_to_Z (wn_spec k) x); intros; auto with zarith.\n";
- fprintf fmt " intros; unfold to_Z, w_0; rewrite (spec_0 (w0_spec)); auto.\n";
- fprintf fmt " Qed.\n";
+ fprintf fmt " Fixpoint safe_shiftl_aux p cont n x {struct p} :=\n";
+ fprintf fmt " safe_shiftl_aux_body \n";
+ fprintf fmt " (fun n x => match p with\n";
+ fprintf fmt " | xH => cont n x\n";
+ fprintf fmt " | xO p => safe_shiftl_aux p (safe_shiftl_aux p cont) n x\n";
+ fprintf fmt " | xI p => safe_shiftl_aux p (safe_shiftl_aux p cont) n x\n";
+ fprintf fmt " end) n x.\n";
fprintf fmt "\n";
- end;
- fprintf fmt " Theorem spec_sub0: forall x y, [x] < [y] -> [sub x y] = 0.\n";
- if gen_proof9 then
+ fprintf fmt " Theorem spec_safe_shift_aux: forall p q n x cont,\n";
+ fprintf fmt " 2 ^ (Zpos q) <= [head0 x] ->\n";
+ fprintf fmt " (forall x, 2 ^ (Zpos p + Zpos q) <= [head0 x] ->\n";
+ fprintf fmt " [cont n x] = [x] * 2 ^ [n]) -> \n";
+ fprintf fmt " [safe_shiftl_aux p cont n x] = [x] * 2 ^ [n].\n";
+ if gen_proof then
begin
fprintf fmt " Proof.\n";
- fprintf fmt " intros x y; case x; unfold sub.\n";
- fprintf fmt " intros x1; case y.\n";
- fprintf fmt " intros y1 H; rewrite spec_w0_sub0; auto.\n";
- for i = 1 to size do
- fprintf fmt " intros y1 H; generalize (spec_w0_eq0 x1); case w0_eq0; intros HH.\n";
- fprintf fmt " unfold to_Z, zero, w_0; rewrite (spec_0 w0_spec); auto.\n";
- if i == 1 then
- fprintf fmt " apply spec_w1_sub0; rewrite spec_extend0_0; auto.\n"
- else
- fprintf fmt " apply spec_w%i_sub0; rewrite spec_extend%i_1; rewrite spec_extend0_0; auto.\n" i (i - 1);
- done;
- fprintf fmt " intros n y1 H; generalize (spec_w0_eq0 x1); case w0_eq0; intros HH.\n";
- fprintf fmt " unfold to_Z, zero, w_0; rewrite (spec_0 w0_spec); auto.\n";
- fprintf fmt " apply spec_wn_sub0; rewrite spec_extendn_0; rewrite spec_extend%i_0; rewrite spec_extend0_0; auto.\n" size;
- for i = 1 to size do
- fprintf fmt " intros x1; case y.\n";
- fprintf fmt " intros y1 H; generalize (spec_w0_eq0 y1); case w0_eq0; intros HH.\n";
- fprintf fmt " generalize H; rewrite HH; unfold to_Z; case (spec_to_Z w%i_spec x1); auto with zarith.\n" i;
- if i = 1 then
- fprintf fmt " apply spec_w1_sub0; rewrite spec_extend0_0; auto.\n"
- else
- fprintf fmt " apply spec_w%i_sub0; rewrite spec_extend%i_1; rewrite spec_extend0_0; auto.\n" i (i-1);
- for j = 1 to size do
- if i <= j then
- fprintf fmt " intros y1 H; apply spec_w%i_sub0; auto.\n" j
- else
- fprintf fmt " intros y1 H; apply spec_w%i_sub0; auto.\n" i;
- done;
- fprintf fmt " intros n y1 H; apply spec_wn_sub0;\n";
- fprintf fmt " rewrite spec_extendn_0; rewrite spec_extend%i_0; auto.\n" (size + 1 - i);
- done;
- fprintf fmt " intros n x1; case y.\n";
- fprintf fmt " intros y1 H; generalize (spec_w0_eq0 y1); case w0_eq0; intros HH.\n";
- fprintf fmt " generalize H; rewrite HH; unfold to_Z; case (spec_to_Z (wn_spec n) x1); auto with zarith.\n";
- fprintf fmt " apply spec_wn_sub0; rewrite spec_extendn_0; \n";
- fprintf fmt " rewrite spec_extend%i_0; rewrite spec_extend0_0; auto.\n" size;
- for i = 1 to size do
- fprintf fmt " intros y1 H; apply spec_wn_sub0; rewrite spec_extendn_0; rewrite spec_extend%i_0; auto.\n" (size + 1 - i);
- done;
- fprintf fmt " intros m y1 H; apply spec_wn_sub0; rewrite spec_cast_l; rewrite spec_cast_r; auto.\n";
- fprintf fmt " Qed.\n"
+ fprintf fmt " intros p; elim p; unfold safe_shiftl_aux; fold safe_shiftl_aux; clear p.\n";
+ fprintf fmt " intros p Hrec q n x cont H1 H2.\n";
+ fprintf fmt " apply spec_safe_shift_aux_body with (q); auto.\n";
+ fprintf fmt " intros x1 H3; apply Hrec with (q + 1)%spositive; auto.\n" "%";
+ fprintf fmt " intros x2 H4; apply Hrec with (p + q + 1)%spositive; auto.\n" "%";
+ fprintf fmt " rewrite <- Pplus_assoc.\n";
+ fprintf fmt " rewrite Zpos_plus_distr; auto.\n";
+ fprintf fmt " intros x3 H5; apply H2.\n";
+ fprintf fmt " rewrite Zpos_xI.\n";
+ fprintf fmt " replace (2 * Zpos p + 1 + Zpos q) with (Zpos p + Zpos (p + q + 1));\n";
+ fprintf fmt " auto.\n";
+ fprintf fmt " repeat rewrite Zpos_plus_distr; ring.\n";
+ fprintf fmt " intros p Hrec q n x cont H1 H2.\n";
+ fprintf fmt " apply spec_safe_shift_aux_body with (q); auto.\n";
+ fprintf fmt " intros x1 H3; apply Hrec with (q); auto.\n";
+ fprintf fmt " apply Zle_trans with (2 := H3); auto with zarith.\n";
+ fprintf fmt " apply Zpower_le_monotone; auto with zarith.\n";
+ fprintf fmt " intros x2 H4; apply Hrec with (p + q)%spositive; auto.\n" "%";
+ fprintf fmt " intros x3 H5; apply H2.\n";
+ fprintf fmt " rewrite (Zpos_xO p).\n";
+ fprintf fmt " replace (2 * Zpos p + Zpos q) with (Zpos p + Zpos (p + q));\n";
+ fprintf fmt " auto.\n";
+ fprintf fmt " repeat rewrite Zpos_plus_distr; ring.\n";
+ fprintf fmt " intros q n x cont H1 H2.\n";
+ fprintf fmt " apply spec_safe_shift_aux_body with (q); auto.\n";
+ fprintf fmt " rewrite Zplus_comm; auto.\n";
+ fprintf fmt " Qed.\n";
end
else
fprintf fmt " Admitted.\n";
fprintf fmt "\n";
- if gen_proof10 then
- begin
-
- fprintf fmt " Fixpoint nmake_op (ww:Set) (ww_op: znz_op ww) (n: nat) : \n";
- fprintf fmt " znz_op (word ww n) :=\n";
- fprintf fmt " match n return znz_op (word ww n) with \n";
- fprintf fmt " O => ww_op\n";
- fprintf fmt " | S n1 => mk_zn2z_op (nmake_op ww ww_op n1) \n";
- fprintf fmt " end.\n";
- fprintf fmt "\n";
- fprintf fmt " Let nmake_op0 := nmake_op _ w0_op.\n";
+ fprintf fmt " Definition safe_shiftl n x :=\n";
+ fprintf fmt " safe_shiftl_aux_body\n";
+ fprintf fmt " (safe_shiftl_aux_body\n";
+ fprintf fmt " (safe_shiftl_aux (digits n) shiftl)) n x.\n";
fprintf fmt "\n";
- fprintf fmt " Theorem nmake_op_S: forall ww (w_op: znz_op ww) x, \n";
- fprintf fmt " nmake_op _ w_op (S x) = mk_zn2z_op (nmake_op _ w_op x).\n";
- fprintf fmt " auto.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt "\n";
- fprintf fmt " Theorem nmake_op0_S: forall x,\n";
- fprintf fmt " nmake_op0 (S x) = mk_zn2z_op (nmake_op0 x).\n";
- fprintf fmt " auto.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt "\n";
- fprintf fmt " Theorem digits_0: znz_digits w0_op = znz_digits (nmake_op0 0).\n";
- fprintf fmt " auto.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt "\n";
- fprintf fmt " Theorem nmake_0: znz_to_Z w0_op = znz_to_Z (nmake_op0 0).\n";
- fprintf fmt " auto.\n";
- fprintf fmt " Qed.\n";
- for i = 1 to size do
- fprintf fmt " Theorem digits_%i: znz_digits w%i_op = znz_digits (nmake_op0 %i).\n" i i i;
- fprintf fmt " rewrite nmake_op0_S; unfold w%i_op.\n" i;
- if i <= 3 then
- fprintf fmt " rewrite digits_zop; rewrite digits_%i.\n" (i - 1)
- else
- fprintf fmt " rewrite digits_kzop; rewrite digits_%i.\n" (i - 1);
- fprintf fmt " rewrite <- digits_zop; auto.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt "\n";
- fprintf fmt " Theorem nmake_%i: znz_to_Z w%i_op = znz_to_Z (nmake_op0 %i).\n" i i i;
- fprintf fmt " rewrite nmake_op0_S; unfold w%i_op.\n" i;
- if i <= 3 then
- fprintf fmt " rewrite make_zop; rewrite digits_%i; rewrite nmake_%i.\n" (i - 1) (i -1)
- else
- fprintf fmt " rewrite make_kzop; rewrite digits_%i; rewrite nmake_%i.\n" (i - 1) (i -1);
- fprintf fmt " rewrite <- make_zop; auto.\n";
- fprintf fmt " Qed.\n";
- done;
-
- fprintf fmt " Let gen_digits: forall n, \n";
- fprintf fmt " base (znz_digits (make_op n)) = (GenBase.gen_wB (znz_digits w%i_op) (S n)).\n" size;
- fprintf fmt " intros n; elim n; clear n.\n";
- fprintf fmt " unfold make_op, make_op_aux; unfold w%i_op; unfold word.\n" (size + 1);
- fprintf fmt " rewrite digits_kzop.\n";
- fprintf fmt " unfold GenBase.gen_wB, GenBase.gen_digits; auto.\n";
-
- fprintf fmt " intros n Hrec; rewrite make_op_S.\n";
- fprintf fmt " change (%sznz_digits (word w%i (S (S n))) (mk_zn2z_op_karatsuba (make_op n))) with (xO (znz_digits (make_op n))).\n" "@" size;
- fprintf fmt " rewrite base_xO; rewrite Hrec.\n";
- fprintf fmt " unfold GenBase.gen_wB; rewrite <- base_xO; auto.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt " \n";
- fprintf fmt " Let gen_make: forall n y, GenBase.gen_to_Z (znz_digits w%i_op) (znz_to_Z w%i_op) (S n) y =\n" size size;
- fprintf fmt " znz_to_Z (make_op n) y.\n";
+ fprintf fmt " Theorem spec_safe_shift: forall n x,\n";
+ fprintf fmt " [safe_shiftl n x] = [x] * 2 ^ [n].\n";
+ if gen_proof then
+ begin
fprintf fmt " Proof.\n";
- fprintf fmt " intros n; elim n; auto.\n";
- fprintf fmt " intros n1 Hrec y; case y; auto.\n";
- fprintf fmt " rewrite make_op_S; auto.\n";
- fprintf fmt " intros yh yl; rewrite znz_to_Z_n.\n";
- fprintf fmt " rewrite gen_digits.\n";
- fprintf fmt " rewrite <- Hrec; rewrite <- Hrec; auto.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt "\n";
-
-
- fprintf fmt " Theorem digits_gend:forall n ww (w_op: znz_op ww), \n";
- fprintf fmt " znz_digits (nmake_op _ w_op n) = \n";
- fprintf fmt " GenBase.gen_digits (znz_digits w_op) n.\n";
- fprintf fmt " Proof.";
- fprintf fmt " intros n; elim n; auto; clear n.\n";
- fprintf fmt " intros n Hrec ww ww_op; simpl GenBase.gen_digits.\n";
- fprintf fmt " rewrite <- Hrec; auto.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt "\n";
- fprintf fmt " Theorem nmake_gen: forall n ww (w_op: znz_op ww), \n";
- fprintf fmt " znz_to_Z (nmake_op _ w_op n) =\n";
- fprintf fmt " %sGenBase.gen_to_Z _ (znz_digits w_op) (znz_to_Z w_op) n.\n" "@";
- fprintf fmt " Proof.";
- fprintf fmt " intros n; elim n; auto; clear n.\n";
- fprintf fmt " intros n Hrec ww ww_op; simpl GenBase.gen_to_Z; unfold zn2z_to_Z.\n";
- fprintf fmt " rewrite <- Hrec; auto.\n";
- fprintf fmt " unfold GenBase.gen_wB; rewrite <- digits_gend; auto.\n";
+ fprintf fmt " intros n x; unfold safe_shiftl, safe_shiftl_aux_body.\n";
+ fprintf fmt " generalize (spec_compare n (head0 x)); case compare; intros H.\n";
+ fprintf fmt " apply spec_shiftl; auto with zarith.\n";
+ fprintf fmt " apply spec_shiftl; auto with zarith.\n";
+ fprintf fmt " rewrite <- (spec_double_size x).\n";
+ fprintf fmt " generalize (spec_compare n (head0 (double_size x))); case compare; intros H1.\n";
+ fprintf fmt " apply spec_shiftl; auto with zarith.\n";
+ fprintf fmt " apply spec_shiftl; auto with zarith.\n";
+ fprintf fmt " rewrite <- (spec_double_size (double_size x)).\n";
+ fprintf fmt " apply spec_safe_shift_aux with 1%spositive.\n" "%";
+ fprintf fmt " apply Zle_trans with (2 := spec_double_size_head0 (double_size x)).\n";
+ fprintf fmt " replace (2 ^ 1) with (2 * 1).\n";
+ fprintf fmt " apply Zmult_le_compat_l; auto with zarith.\n";
+ fprintf fmt " generalize (spec_double_size_head0_pos x); auto with zarith.\n";
+ fprintf fmt " rewrite Zpower_exp_1; ring.\n";
+ fprintf fmt " intros x1 H2; apply spec_shiftl.\n";
+ fprintf fmt " apply Zle_trans with (2 := H2).\n";
+ fprintf fmt " apply Zle_trans with (2 ^ Zpos (digits n)); auto with zarith.\n";
+ fprintf fmt " case (spec_digits n); auto with zarith.\n";
+ fprintf fmt " apply Zpower_le_monotone; auto with zarith.\n";
fprintf fmt " Qed.\n";
+ end
+ else
+ fprintf fmt " Admitted.\n";
fprintf fmt "\n";
-
-
- fprintf fmt " Theorem digits_clean: forall ww (w_op1 w_op2: znz_op ww) n, \n";
- fprintf fmt " znz_digits w_op1 = znz_digits w_op2 ->\n";
- fprintf fmt " znz_digits (nmake_op _ w_op1 n) = znz_digits (nmake_op _ w_op2 n).\n";
- fprintf fmt " Proof.\n";
- fprintf fmt " intros ww w_op1 w_op2 n; elim n; auto; clear n.\n";
- fprintf fmt " intros n Hrec H1.\n";
- fprintf fmt " simpl; unfold zn2z_to_Z; rewrite Hrec; auto.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt " \n";
- fprintf fmt " Theorem nmake_clean: forall ww (w_op1 w_op2: znz_op ww) n, \n";
- fprintf fmt " znz_digits w_op1 = znz_digits w_op2 ->\n";
- fprintf fmt " znz_to_Z w_op1 = znz_to_Z w_op2 ->\n";
- fprintf fmt " znz_to_Z (nmake_op _ w_op1 n) =\n";
- fprintf fmt " znz_to_Z (nmake_op _ w_op2 n).\n";
- fprintf fmt " Proof.\n";
- fprintf fmt " intros ww w_op1 w_op2 n; elim n; auto; clear n.\n";
- fprintf fmt " intros n Hrec H1 H2.\n";
- fprintf fmt " generalize (digits_clean _ _ _ n H1).\n";
- fprintf fmt " simpl; unfold zn2z_to_Z; intros H3.\n";
- fprintf fmt " rewrite Hrec; auto; rewrite H3; auto.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt " \n";
-
- for i = 2 to size do
- for j = 1 to i - 1 do
- fprintf fmt " Let digits_%i_%i: znz_digits (nmake_op _ w%i_op %i) = znz_digits (nmake_op _ w0_op %i).\n" j (i - j) j (i - j) i;
- fprintf fmt " Proof.\n";
- fprintf fmt " replace (nmake_op _ w0_op %i) with (nmake_op _ (nmake_op _ w0_op %i) %i).\n" i j (i - j);
- fprintf fmt " generalize (digits_clean _ _ _ %i digits_%i).\n" (i- j) j;
- fprintf fmt " unfold nmake_op0; auto.\n";
- fprintf fmt " unfold nmake_op; auto.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt "\n";
- done
- done;
-
- for i = 2 to size do
- for j = 1 to i - 1 do
- fprintf fmt " Let nmake_op_%i_%i: znz_to_Z (nmake_op _ w%i_op %i) = znz_to_Z (nmake_op _ w0_op %i).\n" j (i - j) j (i - j) i;
- fprintf fmt " Proof.\n";
- fprintf fmt " replace (nmake_op _ w0_op %i) with (nmake_op _ (nmake_op _ w0_op %i) %i).\n" i j (i - j);
- fprintf fmt " generalize (nmake_clean _ _ _ %i digits_%i nmake_%i).\n" (i- j) j j;
- fprintf fmt " unfold nmake_op0; auto.\n";
- fprintf fmt " unfold nmake_op; auto.\n";
- fprintf fmt " Qed.\n";
- fprintf fmt "\n";
- done
- done
- end;
-
- (* Comparison *)
- fprintf fmt " Theorem spec_compare: forall x y,\n";
- fprintf fmt " match compare x y with \n";
- fprintf fmt " Eq => [x] = [y]\n";
- fprintf fmt " | Lt => [x] < [y]\n";
- fprintf fmt " | Gt => [x] > [y]\n";
- fprintf fmt " end.\n";
- fprintf fmt " Proof.\n";
- if gen_proof11 then
- begin
- for i= 0 to size do
- fprintf fmt " assert(F1_%i:= (spec_0 w%i_spec)).\n" i i;
- fprintf fmt " assert(F2_%i:= (spec_compare w%i_spec (znz_0 w%i_op))).\n" i i i;
- fprintf fmt " assert(F3_%i:= (spec_to_Z w%i_spec)).\n" i i;
- fprintf fmt " assert(F4_%i:= (spec_compare w%i_spec)).\n" i i;
- done;
- fprintf fmt " intros x; case x; clear x; unfold compare, to_Z.\n";
- for i = 0 to size do
- fprintf fmt " intros x y; case y; clear y; auto.\n";
- for j = 0 to i - 1 do
- fprintf fmt " intros y; unfold comparen_%i, w_0, compare_%i.\n" j j;
- fprintf fmt " replace (znz_to_Z w%i_op x) with (%sGenBase.gen_to_Z w%i (znz_digits w%i_op) (znz_to_Z w%i_op) %i x).\n" i "@" j j j (i -j);
- fprintf fmt " apply spec_compare_mn_1; auto.\n";
- fprintf fmt " rewrite <- nmake_gen; rewrite nmake_%i. \n" i;
- if (i == 0) || (j == 0) then
- fprintf fmt " unfold nmake_op0; auto.\n"
- else
- fprintf fmt " rewrite nmake_op_%i_%i; unfold nmake_op0, nmake_op; auto.\n" j (i - j);
- done;
- fprintf fmt " exact (spec_compare w%i_spec x).\n" i;
- for j = i + 1 to size do
- fprintf fmt " intros y; apply spec_opp; unfold comparen_%i, w_0, compare_%i.\n" i i;
- fprintf fmt " replace (znz_to_Z w%i_op y) with (%sGenBase.gen_to_Z w%i (znz_digits w%i_op) (znz_to_Z w%i_op) %i y). \n" j "@" i i i (j - i);
- fprintf fmt " apply spec_compare_mn_1; auto.\n";
- fprintf fmt " rewrite <- nmake_gen; rewrite nmake_%i.\n" j;
- if (i == 0) then
- fprintf fmt " unfold nmake_op0; auto.\n"
- else
- fprintf fmt " rewrite nmake_op_%i_%i; unfold nmake_op0, nmake_op; auto.\n" i (j - i);
- done;
- fprintf fmt " intros n y; apply spec_opp; unfold comparen_%i, w%i, w_0, compare_0.\n" i i;
- fprintf fmt " rewrite <- gen_make.\n";
- fprintf fmt " apply spec_compare_mn_1; auto.\n";
- if i <> size then
- begin
- fprintf fmt " try rewrite (spec_0 w%i_spec); auto.\n" i;
- fprintf fmt " intros x1 y1.\n";
- fprintf fmt " replace (znz_to_Z w%i_op x1) with (%sGenBase.gen_to_Z w%i (znz_digits w%i_op) (znz_to_Z w%i_op) %i x1).\n" size "@" i i i (size - i);
- fprintf fmt " apply spec_compare_mn_1; auto.\n";
- fprintf fmt " rewrite <- nmake_gen; rewrite nmake_%i.\n" size;
- if (i == 0) then
- fprintf fmt " unfold nmake_op0; auto.\n"
- else
- fprintf fmt " rewrite nmake_op_%i_%i; unfold nmake_op0, nmake_op; auto.\n" i (size - i);
- fprintf fmt " intros x1; case (F3_%i x1); split; auto.\n" i;
- fprintf fmt " apply Zlt_trans with (1 := H0); unfold base; apply ZPowerAux.Zpower_lt_monotone.\n";
- fprintf fmt " auto with zarith.\n";
- fprintf fmt " split; [red; intro HH; discriminate HH | idtac].\n";
- fprintf fmt " apply length_pos_lt.\n";
- fprintf fmt " change (length_pos (znz_digits w%i_op)) with\n" size;
- fprintf fmt " (S(%i + (length_pos (znz_digits w%i_op))))%snat.\n" (size - i - 1) i "%";
- fprintf fmt " apply le_lt_n_Sm; apply le_plus_r.\n";
- end;
- done;
- fprintf fmt " intros n x y; case y; clear y.\n";
+ (* even *)
+ fprintf fmt " Definition is_even x :=\n";
+ fprintf fmt " match x with\n";
for i = 0 to size do
- fprintf fmt " intros y; rewrite <- gen_make; unfold comparen_%i; apply spec_compare_mn_1; auto.\n" i;
- if i <> size then
- begin
- fprintf fmt " unfold w_0; try rewrite (spec_0 w%i_spec); auto.\n" i;
- fprintf fmt " intros x1 y1.\n";
- fprintf fmt " replace (znz_to_Z w%i_op x1) with (%sGenBase.gen_to_Z w%i (znz_digits w%i_op) (znz_to_Z w%i_op) %i x1).\n" size "@" i i i (size - i);
- fprintf fmt " apply spec_compare_mn_1; auto.\n";
- fprintf fmt " rewrite <- nmake_gen; rewrite nmake_%i.\n" size;
- if (i == 0) then
- fprintf fmt " unfold nmake_op0; auto.\n"
- else
- fprintf fmt " rewrite nmake_op_%i_%i; unfold nmake_op0, nmake_op; auto.\n" i (size - i);
- fprintf fmt " intros x1; case (F3_%i x1); split; auto.\n" i;
- fprintf fmt " apply Zlt_trans with (1 := H0); unfold base; apply ZPowerAux.Zpower_lt_monotone.\n";
- fprintf fmt " auto with zarith.\n";
- fprintf fmt " split; [red; intro HH; discriminate HH | idtac].\n";
- fprintf fmt " apply length_pos_lt.\n";
- fprintf fmt " change (length_pos (znz_digits w%i_op)) with\n" size;
- fprintf fmt " (S(%i + (length_pos (znz_digits w%i_op))))%snat.\n" (size - i - 1) i "%";
- fprintf fmt " apply le_lt_n_Sm; apply le_plus_r.\n";
- end
+ fprintf fmt " | %s%i wx => w%i_op.(znz_is_even) wx\n" c i i
done;
- fprintf fmt " intros m y.\n";
- fprintf fmt " generalize (spec_cast_l n m x); simpl to_Z; simpl word; intros HH; rewrite <- HH; clear HH.\n";
- fprintf fmt " generalize (spec_cast_r n m y); simpl to_Z; simpl word; intros HH; rewrite <- HH; clear HH.\n";
- fprintf fmt " apply (spec_compare (wn_spec (Max.max n m))).\n";
- fprintf fmt " Qed.\n";
- end else
- fprintf fmt " Admitted.\n";
+ fprintf fmt " | %sn n wx => (make_op n).(znz_is_even) wx\n" c;
+ fprintf fmt " end.\n";
fprintf fmt "\n";
- if gen_proof12 then
+ fprintf fmt " Theorem spec_is_even: forall x,\n";
+ fprintf fmt " if is_even x then [x] mod 2 = 0 else [x] mod 2 = 1.\n";
+ if gen_proof then
begin
- for i = 0 to size do
- fprintf fmt " Let spec_w%i_mul_add: forall x y z,\n" i;
- fprintf fmt " let (q,r) := w%i_mul_add x y z in\n" i;
- fprintf fmt " znz_to_Z w%i_op q * (base (znz_digits w%i_op)) + znz_to_Z w%i_op r =\n" i i i;
- fprintf fmt " znz_to_Z w%i_op x * znz_to_Z w%i_op y + znz_to_Z w%i_op z :=\n" i i i ;
- fprintf fmt " (spec_mul_add w%i_spec).\n" i;
- fprintf fmt "\n";
- done;
- for i = 0 to size do
-
-
- fprintf fmt " Theorem spec_w%i_mul_add_n1: forall n x y z,\n" i;
- fprintf fmt " let (q,r) := w%i_mul_add_n1 n x y z in\n" i;
- fprintf fmt " znz_to_Z w%i_op q * (base (znz_digits (nmake_op _ w%i_op n))) +\n" i i;
- fprintf fmt " znz_to_Z (nmake_op _ w%i_op n) r =\n" i;
- fprintf fmt " znz_to_Z (nmake_op _ w%i_op n) x * znz_to_Z w%i_op y +\n" i i;
- fprintf fmt " znz_to_Z w%i_op z.\n" i;
- fprintf fmt " Proof.\n";
- fprintf fmt " intros n x y z; unfold w%i_mul_add_n1.\n" i;
- fprintf fmt " rewrite nmake_gen.\n";
- fprintf fmt " rewrite digits_gend.\n";
- fprintf fmt " change (base (GenBase.gen_digits (znz_digits w%i_op) n)) with\n" i;
- fprintf fmt " (GenBase.gen_wB (znz_digits w%i_op) n).\n" i;
- fprintf fmt " apply spec_gen_mul_add_n1; auto.\n";
- if i == 0 then fprintf fmt " exact (spec_0 w%i_spec).\n" i;
- fprintf fmt " exact (spec_WW w%i_spec).\n" i;
- fprintf fmt " exact (spec_0W w%i_spec).\n" i;
- fprintf fmt " exact (spec_mul_add w%i_spec).\n" i;
- fprintf fmt " Qed.\n";
- fprintf fmt "\n";
- done;
- end;
-
- fprintf fmt " Theorem spec_mul: forall x y, [mul x y] =[x] * [y].\n";
- fprintf fmt " Proof.\n";
- if gen_proof13 then
- begin
- fprintf fmt " intros x; case x; clear x.\n";
- fprintf fmt " intros x y; case y; clear y; unfold mul.\n";
- fprintf fmt " intros y; rewrite spec_reduce_1; unfold to_Z.\n";
- fprintf fmt " generalize (spec_mul_c w0_spec x y).\n";
- fprintf fmt " intros HH; rewrite <- HH; clear HH; auto.\n";
-
- fprintf fmt " intros y; unfold zero, w0_eq0, w_0.\n";
- fprintf fmt " generalize (spec_eq0 w0_spec x); case znz_eq0.\n";
- fprintf fmt " unfold to_Z; intros HH; rewrite HH; auto.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); auto.\n";
- fprintf fmt " intros _.\n";
- fprintf fmt " generalize (spec_w0_mul_add_n1 1 y x (znz_0 w0_op)); case w0_mul_add_n1.\n";
- fprintf fmt " intros yh yl.\n";
- fprintf fmt " generalize (spec_eq0 w0_spec yh); case znz_eq0.\n";
- fprintf fmt " intros HH; rewrite HH; auto; rewrite Zmult_0_l; rewrite Zplus_0_l; clear HH.\n";
- fprintf fmt " unfold to_Z; rewrite (spec_0 w0_spec); rewrite Zplus_0_r.\n";
- fprintf fmt " rewrite nmake_1; rewrite nmake_0; unfold nmake_op0.\n";
- fprintf fmt " intros HH; rewrite Zmult_comm; rewrite <- HH.\n";
- fprintf fmt " auto.\n";
- fprintf fmt " intros _.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zplus_0_r.\n";
- fprintf fmt " unfold to_Z; rewrite nmake_1; rewrite nmake_0; unfold nmake_op0.\n";
- fprintf fmt " intros HH; rewrite Zmult_comm; rewrite <- HH.\n";
- fprintf fmt " rewrite znz_to_Z_2; rewrite znz_to_Z_1.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zmult_0_l; rewrite Zplus_0_l.\n";
- fprintf fmt " rewrite nmake_1; rewrite nmake_0; rewrite digits_1; unfold nmake_op0; auto.\n";
-
-
- for j = 2 to size - 1 do
- fprintf fmt " intros y; unfold zero, w0_eq0, w_0.\n";
- fprintf fmt " generalize (spec_eq0 w0_spec x); case znz_eq0.\n";
- fprintf fmt " unfold to_Z; intros HH; rewrite HH; auto.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); auto.\n";
- fprintf fmt " intros _.\n";
- fprintf fmt " generalize (spec_w0_mul_add_n1 %i y x (znz_0 w0_op)); case w0_mul_add_n1.\n" j;
- fprintf fmt " intros yh yl.\n";
- fprintf fmt " generalize (spec_eq0 w0_spec yh); case znz_eq0.\n";
- fprintf fmt " intros HH; rewrite HH; auto; rewrite Zmult_0_l; rewrite Zplus_0_l; clear HH.\n";
- fprintf fmt " unfold to_Z; rewrite (spec_0 w0_spec); rewrite Zplus_0_r.\n";
- fprintf fmt " rewrite nmake_%i; rewrite nmake_0; unfold nmake_op0.\n" j;
- fprintf fmt " intros HH; rewrite Zmult_comm; rewrite <- HH.\n";
- fprintf fmt " auto.\n";
- fprintf fmt " intros _.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zplus_0_r.\n";
- fprintf fmt " unfold to_Z; rewrite nmake_%i; rewrite nmake_0; unfold nmake_op0.\n" j;
- fprintf fmt " intros HH; rewrite Zmult_comm; rewrite <- HH; clear HH.\n";
- fprintf fmt " rewrite znz_to_Z_%i.\n" (j + 1);
- fprintf fmt " generalize (spec_extend%i_1 (WW (znz_0 w0_op) yh)); unfold to_Z.\n" (j - 1);
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " rewrite znz_to_Z_1.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zmult_0_l; rewrite Zplus_0_l.\n";
- fprintf fmt " rewrite nmake_%i; rewrite nmake_0; rewrite digits_%i; unfold nmake_op0; auto.\n" j j;
-
- done;
-
- fprintf fmt " intros y; unfold to_Z, zero, w0_eq0, w_0.\n";
- fprintf fmt " generalize (spec_eq0 w0_spec x); case znz_eq0.\n";
- fprintf fmt " intros HH; rewrite HH; auto.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); auto.\n";
- fprintf fmt " intros _.\n";
- fprintf fmt " generalize (spec_w0_mul_add_n1 %i y x (znz_0 w0_op)); case w0_mul_add_n1.\n" size;
- fprintf fmt " intros yh yl.\n";
- fprintf fmt " generalize (spec_eq0 w0_spec yh); case znz_eq0.\n";
- fprintf fmt " intros HH; rewrite HH; auto; rewrite Zmult_0_l; rewrite Zplus_0_l; clear HH.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zplus_0_r.\n";
- fprintf fmt " rewrite nmake_%i; rewrite nmake_0; unfold nmake_op0.\n" size;
- fprintf fmt " intros HH; rewrite Zmult_comm; rewrite <- HH.\n";
- fprintf fmt " auto.\n";
- fprintf fmt " intros _.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zplus_0_r.\n";
- fprintf fmt " rewrite nmake_%i; rewrite nmake_0; unfold nmake_op0.\n" size;
- fprintf fmt " intros HH; rewrite Zmult_comm; rewrite <- HH; clear HH.\n";
- fprintf fmt " simpl make_op; rewrite znz_to_Z_%i.\n" (size + 1);
- fprintf fmt " generalize (spec_extend%i_1 (WW (znz_0 w0_op) yh)); unfold to_Z.\n" (size - 1);
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " rewrite znz_to_Z_1.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zmult_0_l; rewrite Zplus_0_l.\n";
- fprintf fmt " rewrite nmake_%i; rewrite nmake_0; rewrite digits_%i; unfold nmake_op0; auto.\n" size size;
- fprintf fmt " intros n y; unfold to_Z, zero, w0_eq0, w_0.\n";
- fprintf fmt " generalize (spec_eq0 w0_spec x); case znz_eq0.\n";
- fprintf fmt " intros HH; rewrite HH; auto.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); auto.\n";
- fprintf fmt " intros _.\n";
- fprintf fmt " generalize (spec_w%i_mul_add_n1 (S n) y (extend%i w0 (WW (znz_0 w0_op) x)) W0); case w%i_mul_add_n1.\n" size (size - 1) size;
- fprintf fmt " intros yh yl.\n";
- fprintf fmt " rewrite Zplus_0_r.\n";
- fprintf fmt " unfold w%i_eq0; generalize (spec_eq0 w%i_spec yh); case znz_eq0.\n" size size;
- fprintf fmt " intros HH; rewrite HH; auto; rewrite Zmult_0_l; rewrite Zplus_0_l; clear HH.\n";
- fprintf fmt " generalize (spec_extend%i_1 (WW (znz_0 w0_op) x)); unfold to_Z.\n" (size - 1);
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen.\n";
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " rewrite Zmult_comm; rewrite znz_to_Z_1.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zmult_0_l; rewrite Zplus_0_l.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen; auto.\n";
- fprintf fmt " intros _.\n";
- fprintf fmt " rewrite (znz_to_Z_n n).\n";
- fprintf fmt " generalize (spec_extend%i_1 (WW (znz_0 w0_op) x)); unfold to_Z.\n" (size - 1);
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " rewrite znz_to_Z_1.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zmult_0_l; rewrite Zplus_0_l.\n";
- fprintf fmt " generalize (spec_extendn_0 n (WW W0 yh)); unfold to_Z.\n";
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " simpl make_op; rewrite znz_to_Z_%i.\n" (size + 1);
- fprintf fmt " rewrite Zmult_0_l; rewrite Zplus_0_l.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen.\n";
- fprintf fmt " rewrite digits_gend.\n";
- fprintf fmt " rewrite gen_digits; unfold GenBase.gen_wB.\n";
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " rewrite Zmult_comm; auto.\n";
-
- for i = 1 to size do
- fprintf fmt " intros x y; case y; clear y; unfold mul.\n";
- if i = 1 then
- begin
- fprintf fmt " intros y; unfold to_Z, zero, w0_eq0, w_0.\n";
- fprintf fmt " generalize (spec_eq0 w0_spec y); case znz_eq0.\n";
- fprintf fmt " intros HH; rewrite HH; auto; clear HH.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zmult_0_r; auto.\n";
- fprintf fmt " intros _.\n";
- fprintf fmt " generalize (spec_w0_mul_add_n1 1 x y (znz_0 w0_op)); case w0_mul_add_n1.\n";
- fprintf fmt " intros yh yl.\n";
- fprintf fmt " generalize (spec_eq0 w0_spec yh); case znz_eq0.\n";
- fprintf fmt " intros HH; rewrite HH; auto; rewrite Zmult_0_l; rewrite Zplus_0_l; clear HH.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zplus_0_r.\n";
- fprintf fmt " rewrite nmake_1; rewrite nmake_0; unfold nmake_op0.\n";
- fprintf fmt " intros HH; rewrite <- HH.\n";
- fprintf fmt " auto.\n";
- fprintf fmt " intros _.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zplus_0_r.\n";
- fprintf fmt " rewrite nmake_1; rewrite nmake_0; unfold nmake_op0.\n";
- fprintf fmt " intros HH; rewrite <- HH.\n";
- fprintf fmt " rewrite znz_to_Z_2; rewrite znz_to_Z_1.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zmult_0_l; rewrite Zplus_0_l.\n";
- fprintf fmt " rewrite nmake_1; rewrite nmake_0; rewrite digits_1; unfold nmake_op0; auto.\n";
- end
- else if i = size then
- begin
- fprintf fmt " intros y; unfold to_Z, zero, w0_eq0, w_0.\n";
- fprintf fmt " generalize (spec_eq0 w0_spec y); case znz_eq0.\n";
- fprintf fmt " intros HH; rewrite HH; auto; clear HH.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zmult_0_r; auto.\n";
- fprintf fmt " intros _.\n";
- fprintf fmt " generalize (spec_w0_mul_add_n1 %i x y (znz_0 w0_op)); case w0_mul_add_n1.\n" size;
- fprintf fmt " intros yh yl.\n";
- fprintf fmt " generalize (spec_eq0 w0_spec yh); case znz_eq0.\n";
- fprintf fmt " intros HH; rewrite HH; auto; rewrite Zmult_0_l; rewrite Zplus_0_l; clear HH.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zplus_0_r.\n";
- fprintf fmt " rewrite nmake_%i; rewrite nmake_0; unfold nmake_op0.\n" size;
- fprintf fmt " intros HH; rewrite <- HH.\n";
- fprintf fmt " auto.\n";
- fprintf fmt " intros _.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zplus_0_r.\n";
- fprintf fmt " rewrite nmake_%i; rewrite nmake_0; unfold nmake_op0.\n" size;
- fprintf fmt " intros HH; rewrite <- HH; clear HH.\n";
- fprintf fmt " simpl make_op; rewrite znz_to_Z_%i.\n" (size + 1);
- fprintf fmt " generalize (spec_extend%i_1 (WW (znz_0 w0_op) yh)); unfold to_Z.\n" (size - 1);
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " rewrite znz_to_Z_1.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zmult_0_l; rewrite Zplus_0_l.\n";
- fprintf fmt " rewrite nmake_%i; rewrite nmake_0; rewrite digits_%i; unfold nmake_op0; auto.\n" size size;
- end
- else
- begin
- fprintf fmt " intros y; unfold to_Z, zero, w0_eq0, w_0.\n";
- fprintf fmt " generalize (spec_eq0 w0_spec y); case znz_eq0.\n";
- fprintf fmt " intros HH; rewrite HH; auto; clear HH.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zmult_0_r; auto.\n";
- fprintf fmt " intros _.\n";
- fprintf fmt " generalize (spec_w0_mul_add_n1 %i x y (znz_0 w0_op)); case w0_mul_add_n1.\n" i;
- fprintf fmt " intros yh yl.\n";
- fprintf fmt " generalize (spec_eq0 w0_spec yh); case znz_eq0.\n";
- fprintf fmt " intros HH; rewrite HH; auto; rewrite Zmult_0_l; rewrite Zplus_0_l; clear HH.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zplus_0_r.\n";
- fprintf fmt " rewrite nmake_%i; rewrite nmake_0; unfold nmake_op0.\n" i;
- fprintf fmt " intros HH; rewrite <- HH.\n";
- fprintf fmt " auto.\n";
- fprintf fmt " intros _.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zplus_0_r.\n";
- fprintf fmt " rewrite nmake_%i; rewrite nmake_0; unfold nmake_op0.\n" i;
- fprintf fmt " intros HH; rewrite <- HH; clear HH.\n";
- fprintf fmt " rewrite znz_to_Z_%i.\n" (i + 1);
- fprintf fmt " generalize (spec_extend%i_%i (WW (znz_0 w0_op) yh)); unfold to_Z.\n" (i - 1) 1;
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " rewrite znz_to_Z_1.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zmult_0_l; rewrite Zplus_0_l.\n";
- fprintf fmt " rewrite nmake_%i; rewrite nmake_0; rewrite digits_%i; unfold nmake_op0; auto.\n" i i;
- end;
- for j = 1 to size do
- fprintf fmt " (* i = %i j = %i *)\n" i j;
- if i = j then
- begin
- fprintf fmt " intros y; unfold to_Z.\n";
- fprintf fmt " generalize (spec_mul_c w%i_spec x y).\n" i;
- fprintf fmt " intros HH; rewrite <- HH; clear HH; auto.\n";
- end
- else
- begin
- let min,max, wmin, wmax =
- if i < j then i, j, "x", "y" else j, i, "y", "x" in
- if max = size then
- begin
- fprintf fmt " intros y; unfold to_Z, zero, w%i_eq0, w_0.\n" min;
- fprintf fmt " generalize (spec_w%i_mul_add_n1 %i %s %s W0); case w%i_mul_add_n1.\n" min (max -min) wmax wmin min;
- fprintf fmt " intros yh yl.\n";
- fprintf fmt " generalize (spec_eq0 w%i_spec yh); case znz_eq0.\n" min;
- fprintf fmt " intros HH; rewrite HH; auto; rewrite Zmult_0_l; rewrite Zplus_0_l; clear HH.\n";
- fprintf fmt " rewrite Zplus_0_r.\n";
- fprintf fmt " rewrite nmake_%i; rewrite nmake_%i; unfold nmake_op0.\n" max min;
- fprintf fmt " rewrite nmake_op_%i_%i.\n" min (max - min);
- if i = min then
- fprintf fmt " intros HH; rewrite Zmult_comm; rewrite <- HH; clear HH.\n"
- else
- fprintf fmt " intros HH; rewrite <- HH; clear HH.\n";
- fprintf fmt " auto.\n";
- fprintf fmt " intros _.\n";
- fprintf fmt " rewrite Zplus_0_r.\n";
- fprintf fmt " rewrite nmake_%i; rewrite nmake_%i; unfold nmake_op0.\n" max min;
- fprintf fmt " rewrite nmake_op_%i_%i.\n" min (max - min);
- if i = min then
- fprintf fmt " intros HH; rewrite Zmult_comm; rewrite <- HH; clear HH.\n"
- else
- fprintf fmt " intros HH; rewrite <- HH; clear HH.\n";
- fprintf fmt " simpl make_op; rewrite znz_to_Z_%i.\n" (max + 1);
- fprintf fmt " generalize (spec_extend%i_%i yh); unfold to_Z.\n" (max - min) min;
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " rewrite nmake_%i; rewrite digits_%i; unfold nmake_op0.\n" min max;
- fprintf fmt " rewrite nmake_%i; unfold nmake_op0.\n" max;
- fprintf fmt " rewrite digits_%i_%i; auto.\n" min (max - min);
- end
- else
- begin
- fprintf fmt " intros y; unfold to_Z, zero, w0_eq0, w_0.\n";
- fprintf fmt " generalize (spec_w%i_mul_add_n1 %i %s %s W0); case w%i_mul_add_n1.\n" min (max - min) wmax wmin min;
- fprintf fmt " intros yh yl.\n";
- fprintf fmt " unfold w%i_eq0; generalize (spec_eq0 w%i_spec yh); case znz_eq0.\n" min min;
- fprintf fmt " intros HH; rewrite HH; auto; rewrite Zmult_0_l; rewrite Zplus_0_l; clear HH.\n";
- fprintf fmt " rewrite Zplus_0_r.\n";
- fprintf fmt " rewrite nmake_%i; rewrite nmake_%i; unfold nmake_op0.\n" max min;
- if i = min then
- fprintf fmt " rewrite Zmult_comm; rewrite nmake_op_%i_%i; auto.\n" min (max - min)
- else
- fprintf fmt " rewrite nmake_op_%i_%i; auto.\n" min (max - min);
- fprintf fmt " intros _.\n";
- fprintf fmt " rewrite Zplus_0_r.\n";
- fprintf fmt " rewrite digits_%i_%i.\n" min (max - min);
- fprintf fmt " rewrite nmake_%i; rewrite nmake_%i; unfold nmake_op0.\n" min max;
- fprintf fmt " rewrite nmake_op_%i_%i.\n" min (max - min);
- if i = min then
- fprintf fmt " intros HH; rewrite Zmult_comm; rewrite <- HH; clear HH.\n"
- else
- fprintf fmt " intros HH; rewrite <- HH; clear HH.\n";
- fprintf fmt " rewrite znz_to_Z_%i.\n" (max + 1);
- fprintf fmt " generalize (spec_extend%i_%i yh); unfold to_Z.\n" (max - min) min;
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " rewrite nmake_%i; rewrite nmake_%i; rewrite digits_%i; unfold nmake_op0; auto.\n" min max max;
- end
- end
- done;
- if i = size then
- begin
- fprintf fmt " intros n y; unfold to_Z, zero, w%i_eq0, w_0.\n" size;
- fprintf fmt " generalize (spec_w%i_mul_add_n1 (S n) y x W0); case w%i_mul_add_n1.\n" size size;
- fprintf fmt " intros yh yl.\n";
- fprintf fmt " rewrite Zplus_0_r.\n";
- fprintf fmt " generalize (spec_eq0 w%i_spec yh); case znz_eq0.\n" size;
- fprintf fmt " intros HH; rewrite HH; auto; rewrite Zmult_0_l; rewrite Zplus_0_l; clear HH.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen.\n";
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " rewrite Zmult_comm.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen; auto.\n";
- fprintf fmt " intros _.\n";
- fprintf fmt " rewrite (znz_to_Z_n n).\n";
- fprintf fmt " generalize (spec_extendn_0 n (WW W0 yh)); unfold to_Z.\n";
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " simpl make_op; rewrite znz_to_Z_%i.\n" (size + 1);
- fprintf fmt " rewrite Zmult_0_l; rewrite Zplus_0_l.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen.\n";
- fprintf fmt " rewrite digits_gend.\n";
- fprintf fmt " rewrite gen_digits; unfold GenBase.gen_wB.\n";
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " rewrite Zmult_comm; auto.\n";
- end
- else
- begin
- fprintf fmt " intros n y; unfold to_Z, zero, w%i_eq0, w_0.\n" size;
- fprintf fmt " generalize (spec_w%i_mul_add_n1 (S n) y (extend%i _ x) W0); case w%i_mul_add_n1.\n" size (size - i) size;
- fprintf fmt " intros yh yl.\n";
- fprintf fmt " rewrite Zplus_0_r.\n";
- fprintf fmt " unfold w%i_eq0; generalize (spec_eq0 w%i_spec yh); case znz_eq0.\n" size size;
- fprintf fmt " intros HH; rewrite HH; auto; rewrite Zmult_0_l; rewrite Zplus_0_l; clear HH.\n";
- fprintf fmt " generalize (spec_extend%i_%i x); unfold to_Z.\n" (size -i) i;
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen.\n";
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " rewrite Zmult_comm.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen; auto.\n";
- fprintf fmt " intros _.\n";
- fprintf fmt " rewrite (znz_to_Z_n n).\n";
- fprintf fmt " generalize (spec_extend%i_%i x); unfold to_Z.\n" (size - i) i;
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " generalize (spec_extendn_0 n (WW W0 yh)); unfold to_Z.\n";
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " simpl make_op; rewrite znz_to_Z_%i.\n" (size + 1);
- fprintf fmt " rewrite Zmult_0_l; rewrite Zplus_0_l.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen.\n";
- fprintf fmt " rewrite digits_gend.\n";
- fprintf fmt " rewrite gen_digits; unfold GenBase.gen_wB.\n";
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " rewrite Zmult_comm; auto.\n";
- end;
- done;
-
- fprintf fmt " intros n x y; case y; clear y; unfold mul.\n";
- fprintf fmt " intros y; unfold to_Z, zero, w0_eq0, w%i_eq0, w_0.\n" size;
- fprintf fmt " generalize (spec_eq0 w0_spec y); case znz_eq0.\n";
- fprintf fmt " intros HH; rewrite HH; auto; clear HH.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zmult_0_r; auto.\n";
- fprintf fmt " intros _.\n";
- fprintf fmt " generalize (spec_w%i_mul_add_n1 (S n) x (extend%i w0 (WW (znz_0 w0_op) y)) W0); case w%i_mul_add_n1.\n" size (size - 1) size;
- fprintf fmt " intros yh yl.\n";
- fprintf fmt " generalize (spec_eq0 w%i_spec yh); case znz_eq0.\n" size;
- fprintf fmt " intros HH; rewrite HH; auto; rewrite Zmult_0_l; rewrite Zplus_0_l; clear HH.\n";
- fprintf fmt " rewrite Zplus_0_r.\n";
- fprintf fmt " generalize (spec_extend%i_1 (WW (znz_0 w0_op) y)); unfold to_Z.\n" (size - 1);
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " rewrite znz_to_Z_1.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zmult_0_l; rewrite Zplus_0_l.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen; auto.\n";
- fprintf fmt " intros _.\n";
- fprintf fmt " rewrite Zplus_0_r.\n";
- fprintf fmt " generalize (spec_extend%i_1 (WW (znz_0 w0_op) y)); unfold to_Z.\n" (size - 1);
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " rewrite znz_to_Z_1.\n";
- fprintf fmt " rewrite (spec_0 w0_spec); rewrite Zmult_0_l; rewrite Zplus_0_l.\n";
- fprintf fmt " rewrite (znz_to_Z_n n).\n";
- fprintf fmt " generalize (spec_extendn_0 n (WW W0 yh)); unfold to_Z.\n";
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " simpl make_op; rewrite znz_to_Z_%i.\n" (size + 1);
- fprintf fmt " rewrite Zmult_0_l; rewrite Zplus_0_l.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen; auto.\n";
- fprintf fmt " rewrite nmake_%i; rewrite nmake_0; unfold nmake_op0.\n" size;
- fprintf fmt " intros HH; rewrite <- HH; clear HH.\n";
- fprintf fmt " rewrite digits_gend; rewrite gen_digits; auto.\n";
-
- for j = 1 to size - 1 do
- fprintf fmt " intros y; unfold to_Z, zero, w%i_eq0, w_0.\n" size;
- fprintf fmt " generalize (spec_w%i_mul_add_n1 (S n) x (extend%i _ y) W0); case w%i_mul_add_n1.\n" size (size - j) size;
- fprintf fmt " intros yh yl.\n";
- fprintf fmt " generalize (spec_eq0 w%i_spec yh); case znz_eq0.\n" size;
- fprintf fmt " intros HH; rewrite HH; auto; rewrite Zmult_0_l; rewrite Zplus_0_l; clear HH.\n";
- fprintf fmt " rewrite Zplus_0_r.\n";
- fprintf fmt " rewrite nmake_%i; rewrite nmake_%i; unfold nmake_op0.\n" size j;
- fprintf fmt " generalize (spec_extend%i_%i y); unfold to_Z.\n" (size - j) j;
- fprintf fmt " rewrite nmake_%i; unfold nmake_op0.\n" size;
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen.\n";
- fprintf fmt " rewrite nmake_%i; unfold nmake_op0; auto.\n" j;
- fprintf fmt " intros _.\n";
- fprintf fmt " rewrite Zplus_0_r.\n";
- fprintf fmt " generalize (spec_extend%i_%i y); unfold to_Z.\n" (size - j) j;
- fprintf fmt " rewrite nmake_%i; unfold nmake_op0.\n" size;
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " rewrite (znz_to_Z_n n).\n";
- fprintf fmt " generalize (spec_extendn_0 n (WW W0 yh)); unfold to_Z.\n";
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " simpl make_op; rewrite znz_to_Z_%i.\n" (size + 1);
- fprintf fmt " rewrite Zmult_0_l; rewrite Zplus_0_l.\n";
- fprintf fmt " rewrite digits_gend; rewrite gen_digits; auto.\n";
- fprintf fmt " rewrite nmake_%i; unfold nmake_op0; auto.\n" size;
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen; auto.\n";
- done;
- fprintf fmt " intros y; unfold to_Z, zero, w%i_eq0, w_0.\n" size;
- fprintf fmt " generalize (spec_w%i_mul_add_n1 (S n) x y W0); case w%i_mul_add_n1.\n" size size;
- fprintf fmt " intros yh yl.\n";
- fprintf fmt " generalize (spec_eq0 w%i_spec yh); case znz_eq0.\n" size;
- fprintf fmt " intros HH; rewrite HH; auto; rewrite Zmult_0_l; rewrite Zplus_0_l; clear HH.\n";
- fprintf fmt " rewrite Zplus_0_r.\n";
- fprintf fmt " rewrite nmake_%i; unfold nmake_op0.\n" size;
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen; auto.\n";
- fprintf fmt " intros _.\n";
- fprintf fmt " rewrite Zplus_0_r.\n";
- fprintf fmt " rewrite nmake_%i; unfold nmake_op0.\n" size;
- fprintf fmt " rewrite (znz_to_Z_n n).\n";
- fprintf fmt " generalize (spec_extendn_0 n (WW W0 yh)); unfold to_Z.\n";
- fprintf fmt " intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " simpl make_op; rewrite znz_to_Z_%i.\n" (size + 1);
- fprintf fmt " rewrite Zmult_0_l; rewrite Zplus_0_l.\n";
- fprintf fmt " rewrite digits_gend; rewrite gen_digits; auto.\n";
- fprintf fmt " rewrite nmake_%i; unfold nmake_op0; auto.\n" size;
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen.\n";
- fprintf fmt " rewrite <- gen_make; rewrite <- nmake_gen; auto.\n";
- fprintf fmt " intros m y.\n";
- fprintf fmt " rewrite spec_reduce_n.\n";
- fprintf fmt " unfold to_Z.\n";
- fprintf fmt " generalize (spec_mul_c (wn_spec (Max.max n m))\n";
- fprintf fmt " (castm (diff_r n m)\n";
- fprintf fmt " (extend_tr x (snd (diff n m))))\n";
- fprintf fmt " (castm (diff_l n m)\n";
- fprintf fmt " (extend_tr y (fst (diff n m)))));\n";
- fprintf fmt " case znz_mul_c.\n";
- fprintf fmt " generalize (spec_cast_l n m x); simpl to_Z; simpl word; intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " generalize (spec_cast_r n m y); simpl to_Z; simpl word; intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " rewrite make_op_S; auto.\n";
- fprintf fmt " intros x1 y1.\n";
- fprintf fmt " rewrite (znz_to_Z_n (Max.max n m)).\n";
- fprintf fmt " generalize (spec_cast_l n m x); simpl to_Z; simpl word; intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " generalize (spec_cast_r n m y); simpl to_Z; simpl word; intros HH; rewrite HH; clear HH.\n";
- fprintf fmt " simpl zn2z_to_Z; auto.\n";
- fprintf fmt " Qed.\n";
- end
-
- else
- fprintf fmt " Admitted.\n";
- fprintf fmt "\n";
-
- fprintf fmt " Theorem spec_sqrt : forall x,\n";
- fprintf fmt " [sqrt x] ^ 2 <= [x] < ([sqrt x] + 1) ^ 2.\n";
fprintf fmt " Proof.\n";
- if gen_proof14 then
- begin
- fprintf fmt " intros x; case x; unfold sqrt.\n";
+ fprintf fmt " intros x; case x; unfold is_even, to_Z; clear x.\n";
for i = 0 to size do
- fprintf fmt " intros y; rewrite spec_reduce_%i.\n" i;
- fprintf fmt " unfold to_Z, w%i_sqrt.\n" i;
- fprintf fmt " exact (spec_sqrt w%i_spec y).\n" i;
+ fprintf fmt " intros x; exact (spec_is_even w%i_spec x).\n" i;
done;
- fprintf fmt " intros n y; rewrite spec_reduce_n.\n";
- fprintf fmt " exact (spec_sqrt (wn_spec n) y).\n";
- fprintf fmt " Qed.\n";
+ fprintf fmt " intros n x; exact (spec_is_even (wn_spec n) x).\n";
+ fprintf fmt " Qed.\n";
end
else
- fprintf fmt " Admitted.\n";
- fprintf fmt "\n";
+ fprintf fmt " Admitted.\n";
+ fprintf fmt "\n";
+
fprintf fmt "End Make.\n";
@@ -2232,6 +3360,7 @@ let print_Make () =
+
let _ = print_Make ()