aboutsummaryrefslogtreecommitdiff
path: root/src/Curves/Montgomery/Affine.v
blob: bfd7dce60fb45aab2656389a21f217725c450d06 (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
Require Import Crypto.Algebra.Field.
Require Import Crypto.Util.GlobalSettings.
Require Import Crypto.Util.Sum Crypto.Util.Prod.
Require Import Crypto.Util.Tactics.BreakMatch.
Require Import Crypto.Spec.MontgomeryCurve Crypto.Spec.WeierstrassCurve.

Module M.
  Section MontgomeryCurve.
    Import BinNat.
    Context {F Feq Fzero Fone Fopp Fadd Fsub Fmul Finv Fdiv}
            {field:@Algebra.Hierarchy.field F Feq Fzero Fone Fopp Fadd Fsub Fmul Finv Fdiv}
            {Feq_dec:Decidable.DecidableRel Feq}
            {char_ge_3:@Ring.char_ge F Feq Fzero Fone Fopp Fadd Fsub Fmul (BinNat.N.succ_pos (BinNat.N.two))}.
    Local Infix "=" := Feq : type_scope. Local Notation "a <> b" := (not (a = b)) : type_scope.
    Local Infix "+" := Fadd. Local Infix "*" := Fmul.
    Local Infix "-" := Fsub. Local Infix "/" := Fdiv.
    Local Notation "- x" := (Fopp x).
    Local Notation "x ^ 2" := (x*x) (at level 30).
    Local Notation "x ^ 3" := (x*x^2) (at level 30).
    Local Notation "0" := Fzero.  Local Notation "1" := Fone.
    Local Notation "'∞'" := unit : type_scope.
    Local Notation "'∞'" := (inr tt) : core_scope.
    Local Notation "( x , y )" := (inl (pair x y)).
    Local Open Scope core_scope.

    Context {a b: F} {b_nonzero:b <> 0}.

    Program Definition opp (P:@M.point F Feq Fadd Fmul a b) : @M.point F Feq Fadd Fmul a b :=
      match P return F*F+∞ with
      | (x, y) => (x, -y)
      | ∞ => ∞
      end.
    Next Obligation. Proof. destruct P; cbv; break_match; trivial; fsatz. Qed.

    Local Notation add := (M.add(b_nonzero:=b_nonzero)).
    Local Notation point := (@M.point F Feq Fadd Fmul a b).

    Section MontgomeryWeierstrass.
      Local Notation "2" := (1+1).
      Local Notation "3" := (1+2).
      Local Notation "4" := (1+1+1+1).
      Local Notation "9" := (3*3).
      Local Notation "27" := (4*4 + 4+4 +1+1+1).

      Context {char_ge_28:@Ring.char_ge F Feq Fzero Fone Fopp Fadd Fsub Fmul 28}.

      Context {aw bw} {Haw:aw=(3-a^2)/(3*b^2)} {Hbw:bw=(2*a^3-9*a)/(27*b^3)}.
      Local Notation Wpoint := (@W.point F Feq Fadd Fmul aw bw).
      Local Notation Wadd := (@W.add F Feq Fzero Fone Fopp Fadd Fsub Fmul Finv Fdiv field Feq_dec char_ge_3 aw bw).
      Program Definition to_Weierstrass (P:@point) : Wpoint :=
        match M.coordinates P return F*F+∞ with
        | (x, y) => ((x + a/3)/b, y/b)
        | _ => ∞
        end.
      Next Obligation. Proof. destruct P; cbv; break_match; trivial; fsatz. Qed.

      Program Definition of_Weierstrass (P:Wpoint) : point :=
        match W.coordinates P return F*F+∞ with
        | (x,y) => (b*x-a/3, b*y)
        | _ => ∞
        end.
      Next Obligation. Proof. destruct P; cbv; break_match; trivial; fsatz. Qed.
    End MontgomeryWeierstrass.
  End MontgomeryCurve.
End M.