aboutsummaryrefslogtreecommitdiff
path: root/src/Arithmetic/ModularArithmeticTheorems.v
blob: 77186674f791c33cf515488fb349ef9c007f7f10 (plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
230
231
232
233
234
235
236
237
238
239
240
241
242
243
244
245
246
247
248
249
250
251
252
253
254
255
256
257
258
259
260
261
262
263
264
265
266
267
268
269
270
271
272
273
274
275
276
277
278
279
280
281
282
283
284
285
286
287
288
289
290
291
292
293
294
295
296
297
298
299
300
301
302
303
304
305
306
307
308
309
310
311
312
313
314
315
316
317
318
319
320
321
322
323
324
325
326
327
328
329
330
331
332
333
334
335
336
337
338
339
340
341
342
343
344
345
346
Require Import Coq.omega.Omega.
Require Import Crypto.Spec.ModularArithmetic.
Require Import Crypto.Arithmetic.ModularArithmeticPre.

Require Import Coq.ZArith.BinInt Coq.ZArith.Zdiv Coq.ZArith.Znumtheory Coq.NArith.NArith. (* import Zdiv before Znumtheory *)
Require Import Coq.Classes.Morphisms Coq.Setoids.Setoid.
Require Export Coq.setoid_ring.Ring_theory Coq.setoid_ring.Ring_tac.

Require Import Crypto.Algebra.Hierarchy Crypto.Algebra.Ring Crypto.Algebra.Field.
Require Import Crypto.Util.Decidable Crypto.Util.ZUtil.
Require Export Crypto.Util.FixCoqMistakes.

Module F.
  Ltac unwrap_F :=
    intros;
    repeat match goal with [ x : F _ |- _ ] => destruct x end;
    lazy iota beta delta [F.add F.sub F.mul F.opp F.to_Z F.of_Z proj1_sig] in *;
    try apply eqsig_eq;
    pull_Zmod.

  (* FIXME: remove the pose proof once [monoid] no longer contains decidable equality *)
  Global Instance eq_dec {m} : DecidableRel (@eq (F m)). pose proof dec_eq_Z. exact _. Defined.

  Global Instance commutative_ring_modulo m
    : @Algebra.Hierarchy.commutative_ring (F m) Logic.eq 0%F 1%F F.opp F.add F.sub F.mul.
  Proof.
    repeat (split || intro); unwrap_F;
      autorewrite with zsimplify; solve [ exact _ | auto with zarith | congruence].
  Qed.

  Lemma pow_spec {m} a : F.pow a 0%N = 1%F :> F m /\ forall x, F.pow a (1 + x)%N = F.mul a (F.pow a x).
  Proof. change (@F.pow m) with (proj1_sig (@F.pow_with_spec m)); destruct (@F.pow_with_spec m); eauto. Qed.

  Global Instance char_gt {m} :
    @Ring.char_ge
      (F m) Logic.eq F.zero F.one F.opp F.add F.sub F.mul
      m.
  Proof.
  Admitted.

  Section FandZ.
    Context {m:positive}.
    Local Open Scope F_scope.

    Theorem eq_to_Z_iff (x y : F m) : x = y <-> F.to_Z x = F.to_Z y.
    Proof using Type. destruct x, y; intuition; simpl in *; try apply (eqsig_eq _ _); congruence. Qed.

    Lemma eq_of_Z_iff : forall x y : Z, x mod m = y mod m <-> F.of_Z m x = F.of_Z m y.
    Proof using Type. split; unwrap_F; congruence. Qed.


    Lemma to_Z_of_Z : forall z, F.to_Z (F.of_Z m z) = z mod m.
    Proof using Type. unwrap_F; trivial. Qed.

    Lemma of_Z_to_Z x : F.of_Z m (F.to_Z x) = x :> F m.
    Proof using Type. unwrap_F; congruence. Qed.


    Lemma of_Z_mod : forall x, F.of_Z m x = F.of_Z m (x mod m).
    Proof using Type. unwrap_F; trivial. Qed.

    Lemma mod_to_Z : forall (x:F m),  F.to_Z x mod m = F.to_Z x.
    Proof using Type. unwrap_F. congruence. Qed.

    Lemma to_Z_0 : F.to_Z (0:F m) = 0%Z.
    Proof using Type. unwrap_F. apply Zmod_0_l. Qed.

    Lemma of_Z_small_nonzero z : (0 < z < m)%Z -> F.of_Z m z <> 0.
    Proof using Type. intros Hrange Hnz. inversion Hnz. rewrite Zmod_small, Zmod_0_l in *; omega. Qed.

    Lemma to_Z_nonzero (x:F m) : x <> 0 -> F.to_Z x <> 0%Z.
    Proof using Type. intros Hnz Hz. rewrite <- Hz, of_Z_to_Z in Hnz; auto. Qed.

    Lemma to_Z_range (x : F m) : 0 < m -> 0 <= F.to_Z x < m.
    Proof using Type. intros. rewrite <- mod_to_Z. apply Z.mod_pos_bound. trivial. Qed.

    Lemma to_Z_nonzero_range (x : F m) : (x <> 0) -> 0 < m -> (1 <= F.to_Z x < m)%Z.
    Proof using Type.
      unfold not; intros Hnz Hlt.
      rewrite eq_to_Z_iff, to_Z_0 in Hnz; pose proof (to_Z_range x Hlt).
      omega.
    Qed.

    Lemma of_Z_add : forall (x y : Z),
        F.of_Z m (x + y) = F.of_Z m x + F.of_Z m y.
    Proof using Type. unwrap_F; trivial. Qed.

    Lemma to_Z_add : forall x y : F m,
        F.to_Z (x + y) = ((F.to_Z x + F.to_Z y) mod m)%Z.
    Proof using Type. unwrap_F; trivial. Qed.

    Lemma of_Z_mul x y : F.of_Z m (x * y) = F.of_Z _ x * F.of_Z _ y :> F m.
    Proof using Type. unwrap_F. trivial. Qed.

    Lemma to_Z_mul : forall x y : F m,
        F.to_Z (x * y) = ((F.to_Z x * F.to_Z y) mod m)%Z.
    Proof using Type. unwrap_F; trivial. Qed.

    Lemma of_Z_sub x y : F.of_Z _ (x - y) = F.of_Z _ x - F.of_Z _ y :> F m.
    Proof using Type. unwrap_F. trivial. Qed.

    Lemma to_Z_opp : forall x : F m, F.to_Z (F.opp x) = (- F.to_Z x) mod m.
    Proof using Type. unwrap_F; trivial. Qed.

    Lemma of_Z_pow x n : F.of_Z _ x ^ n = F.of_Z _ (x ^ (Z.of_N n) mod m) :> F m.
    Proof using Type.
      induction n as [|n IHn] using N.peano_ind;
        destruct (pow_spec (F.of_Z m x)) as [pow_0 pow_succ] . {
        rewrite pow_0.
        unwrap_F; trivial.
      } {
        rewrite N2Z.inj_succ.
        rewrite Z.pow_succ_r by apply N2Z.is_nonneg.
        rewrite <- N.add_1_l.
        rewrite pow_succ.
        rewrite IHn.
        unwrap_F; trivial.
      }
    Qed.

    Lemma to_Z_pow : forall (x : F m) n,
        F.to_Z (x ^ n)%F = (F.to_Z x ^ Z.of_N n mod m)%Z.
    Proof using Type.
      intros x n.
      symmetry.
      induction n as [|n IHn] using N.peano_ind;
        destruct (pow_spec x) as [pow_0 pow_succ] . {
        rewrite pow_0, Z.pow_0_r; auto.
      } {
        rewrite N2Z.inj_succ.
        rewrite Z.pow_succ_r by apply N2Z.is_nonneg.
        rewrite <- N.add_1_l.
        rewrite pow_succ.
        rewrite <- Zmult_mod_idemp_r.
        rewrite IHn.
        apply to_Z_mul.
      }
    Qed.

    Lemma square_iff (x:F m) :
      (exists y : F m, y * y = x) <-> (exists y : Z, y * y mod m = F.to_Z x)%Z.
    Proof using Type.
      setoid_rewrite eq_to_Z_iff; setoid_rewrite to_Z_mul; split; intro H; destruct H as [x' H].
      - eauto.
      - exists (F.of_Z _ x'); rewrite !to_Z_of_Z; pull_Zmod; auto.
    Qed.
  End FandZ.

  Section FandNat.
    Import NPeano Nat.
    Local Infix "mod" := modulo : nat_scope.
    Local Open Scope nat_scope.

    Context {m:BinPos.positive}.

    Lemma to_nat_of_nat (n:nat) : F.to_nat (F.of_nat m n) = (n mod (Z.to_nat m))%nat.
    Proof using Type.
      unfold F.to_nat, F.of_nat.
      rewrite F.to_Z_of_Z.
      assert (Pos.to_nat m <> 0)%nat as HA by (pose proof Pos2Nat.is_pos m; omega).
      pose proof (mod_Zmod n (Pos.to_nat m) HA) as Hmod.
      rewrite positive_nat_Z in Hmod.
      rewrite <- Hmod.
      rewrite <-Nat2Z.id, Z2Nat.inj_pos; omega.
    Qed.

    Lemma of_nat_to_nat x : F.of_nat m (F.to_nat x) = x.
    Proof using Type.

      unfold F.to_nat, F.of_nat.
      rewrite Z2Nat.id; [ eapply F.of_Z_to_Z | eapply F.to_Z_range; reflexivity].
    Qed.

    Lemma Pos_to_nat_nonzero p : Pos.to_nat p <> 0%nat.
    Admitted.

    Lemma of_nat_mod (n:nat) : F.of_nat m (n mod (Z.to_nat m)) = F.of_nat m n.
    Proof using Type.
      unfold F.of_nat.
      rewrite (F.of_Z_mod (Z.of_nat n)), ?mod_Zmod, ?Z2Nat.id; [reflexivity|..].
      { apply Pos2Z.is_nonneg. }
      { rewrite Z2Nat.inj_pos. apply Pos_to_nat_nonzero. }
    Qed.

    Lemma to_nat_mod (x:F m) (Hm:(0 < m)%Z) : F.to_nat x mod (Z.to_nat m) = F.to_nat x.
    Proof using Type.

      unfold F.to_nat.
      rewrite <-F.mod_to_Z at 2.
      apply Z.mod_to_nat; [assumption|].
      apply F.to_Z_range; assumption.
    Qed.

    Lemma of_nat_add x y :
      F.of_nat m (x + y) = (F.of_nat m x + F.of_nat m y)%F.
    Proof using Type. unfold F.of_nat; rewrite Nat2Z.inj_add, F.of_Z_add; reflexivity. Qed.

    Lemma of_nat_mul x y :
      F.of_nat m (x * y) = (F.of_nat m x * F.of_nat m y)%F.
    Proof using Type. unfold F.of_nat; rewrite Nat2Z.inj_mul, F.of_Z_mul; reflexivity. Qed.
  End FandNat.

  Section RingTacticGadgets.
    Context (m:positive).

    Definition ring_theory : ring_theory 0%F 1%F (@F.add m) (@F.mul m) (@F.sub m) (@F.opp m) eq
      := Algebra.Ring.ring_theory_for_stdlib_tactic.

    Lemma pow_pow_N (x : F m) : forall (n : N), (x ^ id n)%F = pow_N 1%F F.mul x n.
    Proof using Type.
      destruct (pow_spec x) as [HO HS]; intros n.
      destruct n as [|p]; auto; unfold id.
      rewrite ModularArithmeticPre.N_pos_1plus at 1.
      rewrite HS.
      simpl.
      induction p as [|p IHp] using Pos.peano_ind.
      - simpl. rewrite HO. apply Algebra.Hierarchy.right_identity.
      - rewrite (@pow_pos_succ (F m) (@F.mul m) eq _ _ associative x).
        rewrite <-IHp, Pos.pred_N_succ, ModularArithmeticPre.N_pos_1plus, HS.
        trivial.
    Qed.

    Lemma power_theory : power_theory 1%F (@F.mul m) eq id (@F.pow m).
    Proof using Type. split; apply pow_pow_N. Qed.

    (***** Division Theory *****)
    Definition quotrem(a b: F m): F m * F m :=
      let '(q, r) := (Z.quotrem (F.to_Z a) (F.to_Z b)) in (F.of_Z _ q , F.of_Z _ r).
    Lemma div_theory : div_theory eq (@F.add m) (@F.mul m) (@id _) quotrem.
    Proof using Type.
      constructor; intros a b; unfold quotrem, id.

      replace (Z.quotrem (F.to_Z a) (F.to_Z b)) with (Z.quot (F.to_Z a) (F.to_Z b), Z.rem (F.to_Z a) (F.to_Z b)) by
          try (unfold Z.quot, Z.rem; rewrite <- surjective_pairing; trivial).

      unwrap_F; rewrite <-Z.quot_rem'; trivial.
    Qed.

    (* Define a "ring morphism" between GF and Z, i.e. an equivalence
     * between 'inject (ZFunction (X))' and 'GFFunction (inject (X))'.
     *
     * Doing this allows the [ring] tactic to do coefficient
     * manipulations in Z rather than F, because we know it's equivalent
     * to inject the result afterward. *)
    Lemma ring_morph: ring_morph 0%F 1%F F.add F.mul F.sub F.opp   eq
                                 0%Z 1%Z Z.add Z.mul Z.sub Z.opp Z.eqb  (F.of_Z m).
    Proof using Type. split; try intro x; try intro y; unwrap_F; solve [ auto | rewrite (proj1 (Z.eqb_eq x y)); trivial]. Qed.

    (* Redefine our division theory under the ring morphism *)
    Lemma morph_div_theory:
      Ring_theory.div_theory eq Zplus Zmult (F.of_Z m) Z.quotrem.
    Proof using Type.
      split; intros a b.
      replace (Z.quotrem a b) with (Z.quot a b, Z.rem a b);
        try (unfold Z.quot, Z.rem; rewrite <- surjective_pairing; trivial).
      unwrap_F; rewrite <- (Z.quot_rem' a b); trivial.
    Qed.

  End RingTacticGadgets.

  Ltac is_constant t := match t with F.of_Z _ ?x => x | _ => NotConstant end.
  Ltac is_pow_constant t := Ncst t.

  Section VariousModulo.
    Context {m:positive}.
    Local Open Scope F_scope.

    Add Ring _theory : (ring_theory m)
                         (morphism (ring_morph m),
                          constants [is_constant],
                          div (morph_div_theory m),
                          power_tac (power_theory m) [is_pow_constant]).

    Lemma mul_nonzero_l : forall a b : F m, a*b <> 0 -> a <> 0.
    Proof using Type. intros a b Hnz Hz. rewrite Hz in Hnz; apply Hnz; ring. Qed.

    Lemma mul_nonzero_r : forall a b : F m, a*b <> 0 -> b <> 0.
    Proof using Type. intros a b Hnz Hz. rewrite Hz in Hnz; apply Hnz; ring. Qed.
  End VariousModulo.

  Section Pow.
    Context {m:positive}.
    Add Ring _theory' : (ring_theory m)
                          (morphism (ring_morph m),
                           constants [is_constant],
                           div (morph_div_theory m),
                           power_tac (power_theory m) [is_pow_constant]).
    Local Open Scope F_scope.

    Import Algebra.ScalarMult.
    Global Instance pow_is_scalarmult
      : is_scalarmult (G:=F m) (eq:=eq) (add:=F.mul) (zero:=1%F) (mul := fun n x => x ^ (N.of_nat n)).
    Proof using Type.
      split; [ intro P | intros n P | ]; rewrite ?Nat2N.inj_succ, <-?N.add_1_l;
        match goal with
        | [x:F m |- _ ] => solve [destruct (@pow_spec m P); auto]
        | |- Proper _ _ => solve_proper
        end.
    Qed.

    (* TODO: move this somewhere? *)
    Create HintDb nat2N discriminated.
    Hint Rewrite Nat2N.inj_iff
         (eq_refl _ : (0%N = N.of_nat 0))
         (eq_refl _ : (1%N = N.of_nat 1))
         (eq_refl _ : (2%N = N.of_nat 2))
         (eq_refl _ : (3%N = N.of_nat 3))
      : nat2N.
    Hint Rewrite <- Nat2N.inj_double Nat2N.inj_succ_double Nat2N.inj_succ
         Nat2N.inj_add Nat2N.inj_mul Nat2N.inj_sub Nat2N.inj_pred
         Nat2N.inj_div2 Nat2N.inj_max Nat2N.inj_min Nat2N.id
      : nat2N.

    Ltac pow_to_scalarmult_ref :=
      repeat (autorewrite with nat2N;
              match goal with
              | |- context [ (_^?n)%F ] =>
                rewrite <-(N2Nat.id n); generalize (N.to_nat n); clear n;
                intro n
              | |- context [ (_^N.of_nat ?n)%F ] =>
                let rw := constr:(scalarmult_ext(zero:=F.of_Z m 1) n) in
                setoid_rewrite rw (* rewriting moduloa reduction *)
              end).

    Lemma pow_0_r (x:F m) : x^0 = 1.
    Proof using Type. pow_to_scalarmult_ref. apply scalarmult_0_l. Qed.

    Lemma pow_add_r (x:F m) (a b:N) : x^(a+b) = x^a * x^b.
    Proof using Type. pow_to_scalarmult_ref; apply scalarmult_add_l. Qed.

    Lemma pow_0_l (n:N) : n <> 0%N -> 0^n = 0 :> F m.
    Proof using Type. pow_to_scalarmult_ref; destruct n; simpl; intros; [congruence|ring]. Qed.

    Lemma pow_pow_l (x:F m) (a b:N) : (x^a)^b = x^(a*b).
    Proof using Type. pow_to_scalarmult_ref. apply scalarmult_assoc. Qed.

    Lemma pow_1_r (x:F m) : x^1 = x.
    Proof using Type. pow_to_scalarmult_ref; simpl; ring. Qed.

    Lemma pow_2_r (x:F m) : x^2 = x*x.
    Proof using Type. pow_to_scalarmult_ref; simpl; ring. Qed.

    Lemma pow_3_r (x:F m) : x^3 = x*x*x.
    Proof using Type. pow_to_scalarmult_ref; simpl; ring. Qed.
  End Pow.
End F.