aboutsummaryrefslogtreecommitdiff
path: root/src/Encoding
diff options
context:
space:
mode:
authorGravatar jadep <jade.philipoom@gmail.com>2016-04-28 13:13:08 -0400
committerGravatar jadep <jade.philipoom@gmail.com>2016-04-28 13:13:08 -0400
commitcd07805915328fd5ee8d41b6cdd4d0340aa156aa (patch)
tree04a869de660aaa874fca7be13f9fefb86c12cafb /src/Encoding
parent248282849e9b287fe817e64ccf53e09fa3991cbe (diff)
Cleanup: mostly moving lemmas to Util files, some moving lemmas to more general contexts.
Diffstat (limited to 'src/Encoding')
-rw-r--r--src/Encoding/ModularWordEncodingPre.v40
-rw-r--r--src/Encoding/ModularWordEncodingTheorems.v26
-rw-r--r--src/Encoding/PointEncodingPre.v58
3 files changed, 7 insertions, 117 deletions
diff --git a/src/Encoding/ModularWordEncodingPre.v b/src/Encoding/ModularWordEncodingPre.v
index cb2f5a045..417344b43 100644
--- a/src/Encoding/ModularWordEncodingPre.v
+++ b/src/Encoding/ModularWordEncodingPre.v
@@ -3,8 +3,7 @@ Require Import Coq.Numbers.Natural.Peano.NPeano.
Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
Require Import Bedrock.Word.
Require Import Crypto.Tactics.VerdiTactics.
-Require Import Crypto.Util.NatUtil.
-Require Import Crypto.Util.WordUtil.
+Require Import Crypto.Util.ZUtil Crypto.Util.WordUtil.
Require Import Crypto.Spec.Encoding.
Local Open Scope nat_scope.
@@ -21,45 +20,12 @@ Section ModularWordEncodingPre.
else None
.
- (* TODO : move to ZUtil *)
- Lemma bound_check_N : forall x : F m, (Z.to_N x < Npow2 sz)%N.
- Proof.
- intro.
- pose proof (FieldToZ_range x m_pos) as x_range.
- rewrite <- Nnat.N2Nat.id.
- rewrite Npow2_nat.
- replace (Z.to_N x) with (N.of_nat (Z.to_nat x)) by apply Z_nat_N.
- apply (Nat2N_inj_lt _ (pow2 sz)).
- rewrite Zpow_pow2.
- destruct x_range as [x_low x_high].
- apply Z2Nat.inj_lt in x_high; try omega.
- rewrite <- ZUtil.pow_Z2N_Zpow by omega.
- replace (Z.to_nat 2) with 2%nat by auto.
- omega.
- Qed.
-
- (* TODO: move to WordUtil *)
- Lemma wordToN_NToWord_idempotent : forall sz n, (n < Npow2 sz)%N ->
- wordToN (NToWord sz n) = n.
- Proof.
- intros.
- rewrite wordToN_nat, NToWord_nat.
- rewrite wordToNat_natToWord_idempotent; rewrite Nnat.N2Nat.id; auto.
- Qed.
-
- (* TODO: move to WordUtil *)
- Lemma NToWord_wordToN : forall sz w, NToWord sz (wordToN w) = w.
- Proof.
- intros.
- rewrite wordToN_nat, NToWord_nat, Nnat.Nat2N.id.
- apply natToWord_wordToNat.
- Qed.
-
Lemma Fm_encoding_valid : forall x, Fm_dec (Fm_enc x) = Some x.
Proof.
unfold Fm_dec, Fm_enc; intros.
pose proof (FieldToZ_range x m_pos).
- rewrite wordToN_NToWord_idempotent by apply bound_check_N.
+ rewrite wordToN_NToWord_idempotent by (apply bound_check_nat_N;
+ assert (Z.to_nat x < Z.to_nat m) by (apply Z2Nat.inj_lt; omega); omega).
rewrite Z2N.id by omega.
rewrite ZToField_idempotent.
break_if; auto; omega.
diff --git a/src/Encoding/ModularWordEncodingTheorems.v b/src/Encoding/ModularWordEncodingTheorems.v
index 83f199d90..a74ccb522 100644
--- a/src/Encoding/ModularWordEncodingTheorems.v
+++ b/src/Encoding/ModularWordEncodingTheorems.v
@@ -1,10 +1,11 @@
Require Import Coq.ZArith.ZArith Coq.ZArith.Znumtheory.
Require Import Coq.Numbers.Natural.Peano.NPeano.
Require Import Crypto.CompleteEdwardsCurve.CompleteEdwardsCurveTheorems.
-Require Import Crypto.ModularArithmetic.PrimeFieldTheorems.
+Require Import Crypto.ModularArithmetic.PrimeFieldTheorems Crypto.ModularArithmetic.ModularArithmeticTheorems.
Require Import Bedrock.Word.
Require Import Crypto.Tactics.VerdiTactics.
Require Import Crypto.Spec.Encoding.
+Require Import Crypto.Util.ZUtil.
Require Import Crypto.Spec.ModularWordEncoding.
@@ -47,29 +48,6 @@ Section SignBit.
rewrite sign_bit_parity; auto.
Qed.
- (* TODO : move to ZUtil *)
- Lemma Z_odd_mod : forall a b, (b <> 0)%Z ->
- Z.odd (a mod b) = if Z.odd b then xorb (Z.odd a) (Z.odd (a / b)) else Z.odd a.
- Proof.
- intros.
- rewrite Zmod_eq_full by assumption.
- rewrite <-Z.add_opp_r, Z.odd_add, Z.odd_opp, Z.odd_mul.
- break_if; rewrite ?Bool.andb_true_r, ?Bool.andb_false_r; auto using Bool.xorb_false_r.
- Qed.
-
- (* TODO : move to ModularArithmeticTheorems *)
- Lemma F_FieldToZ_add_opp : forall x : F m, x <> 0 -> (FieldToZ x + FieldToZ (opp x) = m)%Z.
- Proof.
- intros.
- rewrite FieldToZ_opp.
- rewrite Z_mod_nz_opp_full, mod_FieldToZ; try omega.
- rewrite mod_FieldToZ.
- replace 0%Z with (@FieldToZ m 0) by auto.
- intro false_eq.
- rewrite <-F_eq in false_eq.
- congruence.
- Qed.
-
Lemma sign_bit_opp : forall (x : F m), x <> 0 -> negb (sign_bit x) = sign_bit (opp x).
Proof.
intros.
diff --git a/src/Encoding/PointEncodingPre.v b/src/Encoding/PointEncodingPre.v
index 9b65d9819..e023da0ea 100644
--- a/src/Encoding/PointEncodingPre.v
+++ b/src/Encoding/PointEncodingPre.v
@@ -118,28 +118,6 @@ Section PointEncoding.
| None => None
end.
- (* TODO : move *)
- Lemma sqrt_mod_q_0 : forall x : F q, sqrt_mod_q x = 0 -> x = 0.
- Proof.
- unfold sqrt_mod_q; intros.
- break_if.
- - match goal with [ H : ?sqrt_x ^ 2 = x, H' : ?sqrt_x = 0 |- _ ] => rewrite <-H, H' end.
- ring.
- - match goal with
- | [H : sqrt_minus1 * _ = 0 |- _ ]=>
- apply Fq_mul_zero_why in H; destruct H as [sqrt_minus1_zero | ? ];
- [ | eapply Fq_root_zero; eauto ]
- end.
- unfold sqrt_minus1 in sqrt_minus1_zero.
- rewrite sqrt_minus1_zero in sqrt_minus1_valid.
- exfalso.
- pose proof (@F_opp_spec q 1) as opp_spec_1.
- rewrite <-sqrt_minus1_valid in opp_spec_1.
- assert (((1 + 0 ^ 2) : F q) = (1 : F q)) as ring_subst by ring.
- rewrite ring_subst in *.
- apply Fq_1_neq_0; assumption.
- Qed.
-
Lemma point_coordinates_encoding_canonical : forall w p,
point_dec_coordinates sign_bit w = Some p -> point_enc_coordinates p = w.
Proof.
@@ -151,7 +129,7 @@ Section PointEncoding.
do 2 (break_if; try congruence); inversion coord_dec_Some; subst.
+ destruct (F_eq_dec (sqrt_mod_q (E.solve_for_x2 f1)) 0%F) as [sqrt_0 | ?].
- rewrite sqrt_0 in *.
- apply sqrt_mod_q_0 in sqrt_0.
+ apply sqrt_mod_q_root_0 in sqrt_0; try assumption.
rewrite sqrt_0 in *.
break_if; [symmetry; auto using Bool.eqb_prop | ].
rewrite sign_bit_zero in *.
@@ -195,19 +173,6 @@ Proof.
exact (encoding_valid y).
Qed.
-(*
-Lemma wordToN_enc_neq_opp : forall x, x <> 0 -> (wordToN (enc (opp x)) <> wordToN (enc x))%N.
-Proof.
- intros x x_nonzero.
- intro false_eq.
- apply x_nonzero.
- apply F_eq_opp_zero; try apply two_lt_q.
- apply wordToN_inj in false_eq.
- apply encoding_inj in false_eq.
- auto.
-Qed.
-*)
-
Lemma sign_bit_opp_eq_iff : forall x y, y <> 0 ->
(sign_bit x <> sign_bit y <-> sign_bit x = sign_bit (opp y)).
Proof.
@@ -242,25 +207,6 @@ Proof.
reflexivity.
Qed.
-(* TODO : move *)
-Lemma sqrt_mod_q_of_0 : @sqrt_mod_q q 0 = 0.
-Proof.
- unfold sqrt_mod_q.
- rewrite !Fq_pow_zero.
- break_if; ring.
-
- congruence.
- intro false_eq.
- SearchAbout Z.to_N.
- rewrite <-(N2Z.id 0) in false_eq.
- rewrite N2Z.inj_0 in false_eq.
- pose proof (prime_ge_2 q prime_q).
- apply Z2N.inj in false_eq; zero_bounds.
- assert (0 < q / 8 + 1)%Z.
- apply Z.add_nonneg_pos; zero_bounds.
- omega.
-Qed.
-
Lemma point_encoding_coordinates_valid : forall p, E.onCurve p ->
point_dec_coordinates sign_bit (point_enc_coordinates p) = Some p.
Proof.
@@ -277,7 +223,7 @@ Proof.
rewrite !Fq_pow_zero, sqrt_mod_q_of_0, Fq_pow_zero by congruence.
rewrite if_F_eq_dec_if_F_eqb, sign_bit_zero.
reflexivity.
- + assert (sqrt_mod_q (x ^ 2) <> 0) by (intro false_eq; apply sqrt_mod_q_0 in false_eq;
+ + assert (sqrt_mod_q (x ^ 2) <> 0) by (intro false_eq; apply sqrt_mod_q_root_0 in false_eq; try assumption;
apply Fq_root_zero in false_eq; rewrite false_eq, F_eqb_refl in eqb_x_0; congruence).
replace (F_eqb (sqrt_mod_q (x ^ 2)) 0) with false by (symmetry;
apply F_eqb_neq_complete; assumption).