Built with Alectryon, running Coq+SerAPI v8.10.0+0.7.0. Coq sources are in this panel; goals and messages will appear in the other. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+🖱️ to focus.
(************************************************************************)
(*         *   The Coq Proof Assistant / The Coq Development Team       *)
(*  v      *   INRIA, CNRS and contributors - Copyright 1999-2018       *)
(* <O___,, *       (see CREDITS file for the list of authors)           *)
(*   \VV/  **************************************************************)
(*    //   *    This file is distributed under the terms of the         *)
(*         *     GNU Lesser General Public License Version 2.1          *)
(*         *     (see LICENSE file for the text of the license)         *)
(************************************************************************)
(*            Benjamin Gregoire, Laurent Thery, INRIA, 2007             *)
(************************************************************************)

Signature and specification of a bounded integer structure

This file specifies how to represent Z/nZ when n=2^d, d being the number of digits of these bounded integers.
Set Implicit Arguments.

Require Import ZArith.
Require Import Znumtheory.
Require Import Zpow_facts.
Require Import DoubleType.

Local Open Scope Z_scope.
First, a description via an operator record and a spec record.
Module ZnZ.

 #[universes(template)]
 Class Ops (t:Type) := MkOps {

    (* Conversion functions with Z *)
    digits : positive;
    zdigits: t;
    to_Z   : t -> Z;
    of_pos : positive -> N * t; (* Euclidean division by [2^digits] *)
    head0  : t -> t; (* number of digits 0 in front of the number *)
    tail0  : t -> t; (* number of digits 0 at the bottom of the number *)

    (* Basic numbers *)
    zero : t;
    one  : t;
    minus_one : t;  (* [2^digits-1], which is equivalent to [-1] *)

    (* Comparison *)
    compare     : t -> t -> comparison;
    eq0         : t -> bool;

    (* Basic arithmetic operations *)
    opp_c       : t -> carry t;
    opp         : t -> t;
    opp_carry   : t -> t; (* the carry is known to be -1 *)

    succ_c      : t -> carry t;
    add_c       : t -> t -> carry t;
    add_carry_c : t -> t -> carry t;
    succ        : t -> t;
    add         : t -> t -> t;
    add_carry   : t -> t -> t;

    pred_c      : t -> carry t;
    sub_c       : t -> t -> carry t;
    sub_carry_c : t -> t -> carry t;
    pred        : t -> t;
    sub         : t -> t -> t;
    sub_carry   : t -> t -> t;

    mul_c       : t -> t -> zn2z t;
    mul         : t -> t -> t;
    square_c    : t -> zn2z t;

    (* Special divisions operations *)
    div21       : t -> t -> t -> t*t;
    div_gt      : t -> t -> t * t; (* specialized version of [div] *)
    div         : t -> t -> t * t;

    modulo_gt   : t -> t -> t; (* specialized version of [mod] *)
    modulo      : t -> t -> t;

    gcd_gt      : t -> t -> t; (* specialized version of [gcd] *)
    gcd         : t -> t -> t;
    (* [add_mul_div p i j] is a combination of the [(digits-p)]
       low bits of [i] above the [p] high bits of [j]:
       [add_mul_div p i j = i*2^p+j/2^(digits-p)] *)
    add_mul_div : t -> t -> t -> t;
    (* [pos_mod p i] is [i mod 2^p] *)
    pos_mod     : t -> t -> t;

    is_even     : t -> bool;
    (* square root *)
    sqrt2       : t -> t -> t * carry t;
    sqrt        : t -> t;
    (* bitwise operations *)
    lor         : t -> t -> t;
    land        : t -> t -> t;
    lxor        : t -> t -> t }.
 
 Section Specs.
 Context {t : Type}{ops : Ops t}.

 Notation "[| x |]" := (to_Z x)  (at level 0, x at level 99).

 Let wB := base digits.

 Notation "[+| c |]" :=
   (interp_carry 1 wB to_Z c)  (at level 0, c at level 99).

 Notation "[-| c |]" :=
   (interp_carry (-1) wB to_Z c)  (at level 0, c at level 99).

 Notation "[|| x ||]" :=
   (zn2z_to_Z wB to_Z x)  (at level 0, x at level 99).

 Class Specs := MkSpecs {

    (* Conversion functions with Z *)
    spec_to_Z   : forall x, 0 <= [| x |] < wB;
    spec_of_pos : forall p,
           Zpos p = (Z.of_N (fst (of_pos p)))*wB + [|(snd (of_pos p))|];
    spec_zdigits : [| zdigits |] = Zpos digits;
    spec_more_than_1_digit: 1 < Zpos digits;

    (* Basic numbers *)
    spec_0   : [|zero|] = 0;
    spec_1   : [|one|] = 1;
    spec_m1 : [|minus_one|] = wB - 1;

    (* Comparison *)
    spec_compare : forall x y, compare x y = ([|x|] ?= [|y|]);
    (* NB: the spec of [eq0] is deliberately partial,
       see DoubleCyclic where [eq0 x = true <-> x = W0] *)
    spec_eq0 : forall x, eq0 x = true -> [|x|] = 0;
    (* Basic arithmetic operations *)
    spec_opp_c : forall x, [-|opp_c x|] = -[|x|];
    spec_opp : forall x, [|opp x|] = (-[|x|]) mod wB;
    spec_opp_carry : forall x, [|opp_carry x|] = wB - [|x|] - 1;

    spec_succ_c : forall x, [+|succ_c x|] = [|x|] + 1;
    spec_add_c  : forall x y, [+|add_c x y|] = [|x|] + [|y|];
    spec_add_carry_c : forall x y, [+|add_carry_c x y|] = [|x|] + [|y|] + 1;
    spec_succ : forall x, [|succ x|] = ([|x|] + 1) mod wB;
    spec_add : forall x y, [|add x y|] = ([|x|] + [|y|]) mod wB;
    spec_add_carry :
	 forall x y, [|add_carry x y|] = ([|x|] + [|y|] + 1) mod wB;

    spec_pred_c : forall x, [-|pred_c x|] = [|x|] - 1;
    spec_sub_c : forall x y, [-|sub_c x y|] = [|x|] - [|y|];
    spec_sub_carry_c : forall x y, [-|sub_carry_c x y|] = [|x|] - [|y|] - 1;
    spec_pred : forall x, [|pred x|] = ([|x|] - 1) mod wB;
    spec_sub : forall x y, [|sub x y|] = ([|x|] - [|y|]) mod wB;
    spec_sub_carry :
	 forall x y, [|sub_carry x y|] = ([|x|] - [|y|] - 1) mod wB;

    spec_mul_c : forall x y, [|| mul_c x y ||] = [|x|] * [|y|];
    spec_mul : forall x y, [|mul x y|] = ([|x|] * [|y|]) mod wB;
    spec_square_c : forall x, [|| square_c x||] = [|x|] * [|x|];

    (* Special divisions operations *)
    spec_div21 : forall a1 a2 b,
      wB/2 <= [|b|] ->
      [|a1|] < [|b|] ->
      let (q,r) := div21 a1 a2 b in
      [|a1|] *wB+ [|a2|] = [|q|] * [|b|] + [|r|] /\
      0 <= [|r|] < [|b|];
    spec_div_gt : forall a b, [|a|] > [|b|] -> 0 < [|b|] ->
      let (q,r) := div_gt a b in
      [|a|] = [|q|] * [|b|] + [|r|] /\
      0 <= [|r|] < [|b|];
    spec_div : forall a b, 0 < [|b|] ->
      let (q,r) := div a b in
      [|a|] = [|q|] * [|b|] + [|r|] /\
      0 <= [|r|] < [|b|];

    spec_modulo_gt : forall a b, [|a|] > [|b|] -> 0 < [|b|] ->
      [|modulo_gt a b|] = [|a|] mod [|b|];
    spec_modulo :  forall a b, 0 < [|b|] ->
      [|modulo a b|] = [|a|] mod [|b|];

    spec_gcd_gt : forall a b, [|a|] > [|b|] ->
      Zis_gcd [|a|] [|b|] [|gcd_gt a b|];
    spec_gcd : forall a b, Zis_gcd [|a|] [|b|] [|gcd a b|];


    (* shift operations *)
    spec_head00:  forall x, [|x|] = 0 -> [|head0 x|] = Zpos digits;
    spec_head0  : forall x,  0 < [|x|] ->
	 wB/ 2 <= 2 ^ ([|head0 x|]) * [|x|] < wB;
    spec_tail00:  forall x, [|x|] = 0 -> [|tail0 x|] = Zpos digits;
    spec_tail0  : forall x, 0 < [|x|] ->
         exists y, 0 <= y /\ [|x|] = (2 * y + 1) * (2 ^ [|tail0 x|]) ;
    spec_add_mul_div : forall x y p,
       [|p|] <= Zpos digits ->
       [| add_mul_div p x y |] =
         ([|x|] * (2 ^ [|p|]) +
          [|y|] / (2 ^ ((Zpos digits) - [|p|]))) mod wB;
    spec_pos_mod : forall w p,
       [|pos_mod p w|] = [|w|] mod (2 ^ [|p|]);
    (* sqrt *)
    spec_is_even : forall x,
      if is_even x then [|x|] mod 2 = 0 else [|x|] mod 2 = 1;
    spec_sqrt2 : forall x y,
       wB/ 4 <= [|x|] ->
       let (s,r) := sqrt2 x y in
          [||WW x y||] = [|s|] ^ 2 + [+|r|] /\
          [+|r|] <= 2 * [|s|];
    spec_sqrt : forall x,
       [|sqrt x|] ^ 2 <= [|x|] < ([|sqrt x|] + 1) ^ 2;
    spec_lor : forall x y, [|lor x y|] = Z.lor [|x|] [|y|];
    spec_land : forall x y, [|land x y|] = Z.land [|x|] [|y|];
    spec_lxor : forall x y, [|lxor x y|] = Z.lxor [|x|] [|y|]
  }.

 End Specs.

 Arguments Specs {t} ops.
Generic construction of double words
 Section WW.

 Context {t : Type}{ops : Ops t}{specs : Specs ops}.

 Let wB := base digits.

 Definition WO' (eq0:t->bool) zero h :=
  if eq0 h then W0 else WW h zero.

 Definition WO := Eval lazy beta delta [WO'] in
   let eq0 := ZnZ.eq0 in
   let zero := ZnZ.zero in
   WO' eq0 zero.

 Definition OW' (eq0:t->bool) zero l :=
  if eq0 l then W0 else WW zero l.

 Definition OW := Eval lazy beta delta [OW'] in
   let eq0 := ZnZ.eq0 in
   let zero := ZnZ.zero in
   OW' eq0 zero.

 Definition WW'  (eq0:t->bool) zero h l :=
  if eq0 h then OW' eq0 zero l else WW h l.

 Definition WW := Eval lazy beta delta [WW' OW'] in
   let eq0 := ZnZ.eq0 in
   let zero := ZnZ.zero in
   WW' eq0 zero.

 
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z

forall h : t, zn2z_to_Z wB to_Z (WO h) = to_Z h * wB
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z

forall h : t, zn2z_to_Z wB to_Z (WO h) = to_Z h * wB
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z
h:t

match (if eq0 h then W0 else DoubleType.WW h zero) with | W0 => 0 | DoubleType.WW xh xl => to_Z xh * wB + to_Z xl end = to_Z h * wB
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z
h:t
H:eq0 h = true

0 = to_Z h * wB
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z
h:t
H:eq0 h = false
to_Z h * wB + to_Z zero = to_Z h * wB
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z
h:t
H:eq0 h = false

to_Z h * wB + to_Z zero = to_Z h * wB
rewrite spec_0; auto with zarith. Qed.
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z

forall l : t, zn2z_to_Z wB to_Z (OW l) = to_Z l
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z

forall l : t, zn2z_to_Z wB to_Z (OW l) = to_Z l
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z
l:t

match (if eq0 l then W0 else DoubleType.WW zero l) with | W0 => 0 | DoubleType.WW xh xl => to_Z xh * wB + to_Z xl end = to_Z l
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z
l:t
H:eq0 l = true

0 = to_Z l
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z
l:t
H:eq0 l = false
to_Z zero * wB + to_Z l = to_Z l
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z
l:t
H:eq0 l = false

to_Z zero * wB + to_Z l = to_Z l
rewrite spec_0; auto with zarith. Qed.
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z

forall h l : t, zn2z_to_Z wB to_Z (WW h l) = to_Z h * wB + to_Z l
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z

forall h l : t, zn2z_to_Z wB to_Z (WW h l) = to_Z h * wB + to_Z l
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z
h, l:t

zn2z_to_Z wB to_Z (if eq0 h then if eq0 l then W0 else DoubleType.WW zero l else DoubleType.WW h l) = to_Z h * wB + to_Z l
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z
h, l:t
H:eq0 h = true

zn2z_to_Z wB to_Z (if eq0 l then W0 else DoubleType.WW zero l) = to_Z h * wB + to_Z l
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z
h, l:t
H:eq0 h = false
zn2z_to_Z wB to_Z (DoubleType.WW h l) = to_Z h * wB + to_Z l
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z
h, l:t
H:eq0 h = true

zn2z_to_Z wB to_Z (if eq0 l then W0 else DoubleType.WW zero l) = 0 * wB + to_Z l
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z
h, l:t
H:eq0 h = false
zn2z_to_Z wB to_Z (DoubleType.WW h l) = to_Z h * wB + to_Z l
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z
h, l:t
H:eq0 h = true

zn2z_to_Z wB to_Z (OW l) = 0 * wB + to_Z l
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z
h, l:t
H:eq0 h = false
zn2z_to_Z wB to_Z (DoubleType.WW h l) = to_Z h * wB + to_Z l
t:Type
ops:Ops t
specs:Specs ops
wB:=base digits:Z
h, l:t
H:eq0 h = false

zn2z_to_Z wB to_Z (DoubleType.WW h l) = to_Z h * wB + to_Z l
simpl; auto. Qed. End WW.
Injecting Z numbers into a cyclic structure
 Section Of_Z.

 Context {t : Type}{ops : Ops t}{specs : Specs ops}.

 Notation "[| x |]" := (to_Z x)  (at level 0, x at level 99).

 
t:Type
ops:Ops t
specs:Specs ops

forall p : positive, Z.pos p < base digits -> [|snd (of_pos p)|] = Z.pos p
t:Type
ops:Ops t
specs:Specs ops

forall p : positive, Z.pos p < base digits -> [|snd (of_pos p)|] = Z.pos p
t:Type
ops:Ops t
specs:Specs ops
p:positive
Hp:Z.pos p < base digits

[|snd (of_pos p)|] = Z.pos p
t:Type
ops:Ops t
specs:Specs ops
p:positive
Hp:Z.pos p < base digits

Z.pos p = Z.of_N (fst (of_pos p)) * base digits + [|snd (of_pos p)|] -> [|snd (of_pos p)|] = Z.pos p
t:Type
ops:Ops t
specs:Specs ops
p:positive
Hp:Z.pos p < base digits
n:N
w1:t

Z.pos p = Z.of_N n * base digits + [|w1|] -> [|w1|] = Z.pos p
t:Type
ops:Ops t
specs:Specs ops
p:positive
Hp:Z.pos p < base digits
n:N
w1:t

forall p0 : positive, Z.pos p = Z.of_N (N.pos p0) * base digits + [|w1|] -> [|w1|] = Z.pos p
t:Type
ops:Ops t
specs:Specs ops
p:positive
n:N
w1:t
p1:positive
Hp1:Z.pos p = Z.of_N (N.pos p1) * base digits + [|w1|]

base digits <= Z.pos p
t:Type
ops:Ops t
specs:Specs ops
p:positive
n:N
w1:t
p1:positive
Hp1:Z.pos p = Z.of_N (N.pos p1) * base digits + [|w1|]

1 * base digits + 0 <= Z.pos p
t:Type
ops:Ops t
specs:Specs ops
p:positive
n:N
w1:t
p1:positive
Hp1:Z.pos p = Z.of_N (N.pos p1) * base digits + [|w1|]

1 * base digits + 0 <= Z.of_N (N.pos p1) * base digits + [|w1|]
t:Type
ops:Ops t
specs:Specs ops
p:positive
n:N
w1:t
p1:positive
Hp1:Z.pos p = Z.of_N (N.pos p1) * base digits + [|w1|]

1 * base digits <= Z.of_N (N.pos p1) * base digits
t:Type
ops:Ops t
specs:Specs ops
p:positive
n:N
w1:t
p1:positive
Hp1:Z.pos p = Z.of_N (N.pos p1) * base digits + [|w1|]
0 <= [|w1|]
t:Type
ops:Ops t
specs:Specs ops
p:positive
n:N
w1:t
p1:positive
Hp1:Z.pos p = Z.of_N (N.pos p1) * base digits + [|w1|]

1 <= Z.of_N (N.pos p1)
t:Type
ops:Ops t
specs:Specs ops
p:positive
n:N
w1:t
p1:positive
Hp1:Z.pos p = Z.of_N (N.pos p1) * base digits + [|w1|]
0 <= base digits
t:Type
ops:Ops t
specs:Specs ops
p:positive
n:N
w1:t
p1:positive
Hp1:Z.pos p = Z.of_N (N.pos p1) * base digits + [|w1|]
0 <= [|w1|]
t:Type
ops:Ops t
specs:Specs ops
p:positive
n:N
w1:t
p1:positive
Hp1:Z.pos p = Z.of_N (N.pos p1) * base digits + [|w1|]

0 <= base digits
t:Type
ops:Ops t
specs:Specs ops
p:positive
n:N
w1:t
p1:positive
Hp1:Z.pos p = Z.of_N (N.pos p1) * base digits + [|w1|]
0 <= [|w1|]
t:Type
ops:Ops t
specs:Specs ops
p:positive
n:N
w1:t
p1:positive
Hp1:Z.pos p = Z.of_N (N.pos p1) * base digits + [|w1|]

0 <= [|w1|]
case (spec_to_Z w1); auto with zarith. Qed. Definition of_Z z := match z with | Zpos p => snd (of_pos p) | _ => zero end.
t:Type
ops:Ops t
specs:Specs ops

forall p : Z, 0 <= p < base digits -> [|of_Z p|] = p
t:Type
ops:Ops t
specs:Specs ops

forall p : Z, 0 <= p < base digits -> [|of_Z p|] = p
t:Type
ops:Ops t
specs:Specs ops
p:Z

forall p0 : positive, 0 <= Z.pos p0 < base digits -> [|snd (of_pos p0)|] = Z.pos p0
t:Type
ops:Ops t
specs:Specs ops
p:Z
forall p0 : positive, 0 <= Z.neg p0 < base digits -> 0 = Z.neg p0
t:Type
ops:Ops t
specs:Specs ops
p:Z

forall p0 : positive, 0 <= Z.neg p0 < base digits -> 0 = Z.neg p0
intros p1 (H1, _); contradict H1; apply Z.lt_nge; red; simpl; auto. Qed. End Of_Z. End ZnZ.
A modular specification grouping the earlier records.
Module Type CyclicType.
 Parameter t : Type.
 Declare Instance ops : ZnZ.Ops t.
 Declare Instance specs : ZnZ.Specs ops.
End CyclicType.
A Cyclic structure can be seen as a ring
Module CyclicRing (Import Cyclic : CyclicType).

Local Notation "[| x |]" := (ZnZ.to_Z x) (at level 0, x at level 99).

Definition eq (n m : t) := [| n |] = [| m |].

Local Infix "=="  := eq (at level 70).
Local Notation "0" := ZnZ.zero.
Local Notation "1" := ZnZ.one.
Local Infix "+" := ZnZ.add.
Local Infix "-" := ZnZ.sub.
Local Notation "- x" := (ZnZ.opp x).
Local Infix "*" := ZnZ.mul.
Notation wB := (base ZnZ.digits).

Hint Rewrite ZnZ.spec_0 ZnZ.spec_1 ZnZ.spec_add ZnZ.spec_mul
 ZnZ.spec_opp ZnZ.spec_sub
 : cyclic.

Ltac zify := unfold eq in *; autorewrite with cyclic.


forall x : t, 0 + x == x

forall x : t, 0 + x == x
x:t

0 + x == x
x:t

(0 + [|x|]) mod wB = [|x|]
x:t

[|x|] mod wB = [|x|]
x:t

0 <= [|x|] < wB
apply ZnZ.spec_to_Z. Qed.

forall x y : t, x + y == y + x

forall x y : t, x + y == y + x
x, y:t

x + y == y + x
x, y:t

([|x|] + [|y|]) mod wB = ([|y|] + [|x|]) mod wB
now rewrite Z.add_comm. Qed.

forall x y z : t, x + (y + z) == x + y + z

forall x y z : t, x + (y + z) == x + y + z
x, y, z:t

x + (y + z) == x + y + z
x, y, z:t

([|x|] + ([|y|] + [|z|]) mod wB) mod wB = (([|x|] + [|y|]) mod wB + [|z|]) mod wB
now rewrite Zplus_mod_idemp_r, Zplus_mod_idemp_l, Z.add_assoc. Qed.

forall x : t, 1 * x == x

forall x : t, 1 * x == x
x:t

1 * x == x
x:t

(1 * [|x|]) mod wB = [|x|]
x:t

[|x|] mod wB = [|x|]
x:t

0 <= [|x|] < wB
apply ZnZ.spec_to_Z. Qed.

forall x y : t, x * y == y * x

forall x y : t, x * y == y * x
x, y:t

x * y == y * x
x, y:t

([|x|] * [|y|]) mod wB = ([|y|] * [|x|]) mod wB
now rewrite Z.mul_comm. Qed.

forall x y z : t, x * (y * z) == x * y * z

forall x y z : t, x * (y * z) == x * y * z
x, y, z:t

x * (y * z) == x * y * z
x, y, z:t

([|x|] * (([|y|] * [|z|]) mod wB)) mod wB = (([|x|] * [|y|]) mod wB * [|z|]) mod wB
now rewrite Zmult_mod_idemp_r, Zmult_mod_idemp_l, Z.mul_assoc. Qed.

forall x y z : t, (x + y) * z == x * z + y * z

forall x y z : t, (x + y) * z == x * z + y * z
x, y, z:t

(x + y) * z == x * z + y * z
x, y, z:t

(([|x|] + [|y|]) mod wB * [|z|]) mod wB = (([|x|] * [|z|]) mod wB + ([|y|] * [|z|]) mod wB) mod wB
now rewrite <- Zplus_mod, Zmult_mod_idemp_l, Z.mul_add_distr_r. Qed.

forall x y : t, x + - y == x - y

forall x y : t, x + - y == x - y
x, y:t

x + - y == x - y
x, y:t

([|x|] + - [|y|] mod wB) mod wB = ([|x|] - [|y|]) mod wB
x, y:t

([|x|] + - [|y|] mod wB) mod wB = ([|x|] - [|y|] mod wB) mod wB
x, y:t

([|x|] + - [|y|] mod wB) mod wB = ([|x|] + - ([|y|] mod wB)) mod wB
x, y:t
EQ:[|y|] mod wB = 0

([|x|] + - [|y|] mod wB) mod wB = ([|x|] + - ([|y|] mod wB)) mod wB
x, y:t
NEQ:[|y|] mod wB <> 0
([|x|] + - [|y|] mod wB) mod wB = ([|x|] + - ([|y|] mod wB)) mod wB
x, y:t
NEQ:[|y|] mod wB <> 0

([|x|] + - [|y|] mod wB) mod wB = ([|x|] + - ([|y|] mod wB)) mod wB
x, y:t
NEQ:[|y|] mod wB <> 0

([|x|] + (wB - [|y|] mod wB)) mod wB = ([|x|] + - ([|y|] mod wB)) mod wB
x, y:t
NEQ:[|y|] mod wB <> 0

([|x|] + (wB mod wB - [|y|] mod wB) mod wB) mod wB = ([|x|] + - ([|y|] mod wB)) mod wB
x, y:t
NEQ:[|y|] mod wB <> 0

([|x|] + (0 - [|y|] mod wB) mod wB) mod wB = ([|x|] + - ([|y|] mod wB)) mod wB
x, y:t
NEQ:[|y|] mod wB <> 0

([|x|] + - ([|y|] mod wB) mod wB) mod wB = ([|x|] + - ([|y|] mod wB)) mod wB
now rewrite Zplus_mod_idemp_r. Qed.

forall x : t, x + - x == 0

forall x : t, x + - x == 0
x:t

x + - x == 0
x:t

[|x + - x|] = [|0|]
x:t

[|x - x|] = [|0|]
x:t

([|x|] - [|x|]) mod wB = 0
now rewrite Z.sub_diag, Zmod_0_l. Qed.

ring_theory 0 1 ZnZ.add ZnZ.mul ZnZ.sub ZnZ.opp eq

ring_theory 0 1 ZnZ.add ZnZ.mul ZnZ.sub ZnZ.opp eq

forall x : t, 0 + x == x

forall x y : t, x + y == y + x

forall x y z : t, x + (y + z) == x + y + z

forall x : t, 1 * x == x

forall x y : t, x * y == y * x

forall x y z : t, x * (y * z) == x * y * z

forall x y z : t, (x + y) * z == x * z + y * z

forall x y : t, x - y == x + - y

forall x : t, x + - x == 0

forall x y : t, x + y == y + x

forall x y z : t, x + (y + z) == x + y + z

forall x : t, 1 * x == x

forall x y : t, x * y == y * x

forall x y z : t, x * (y * z) == x * y * z

forall x y z : t, (x + y) * z == x * z + y * z

forall x y : t, x - y == x + - y

forall x : t, x + - x == 0

forall x y z : t, x + (y + z) == x + y + z

forall x : t, 1 * x == x

forall x y : t, x * y == y * x

forall x y z : t, x * (y * z) == x * y * z

forall x y z : t, (x + y) * z == x * z + y * z

forall x y : t, x - y == x + - y

forall x : t, x + - x == 0

forall x : t, 1 * x == x

forall x y : t, x * y == y * x

forall x y z : t, x * (y * z) == x * y * z

forall x y z : t, (x + y) * z == x * z + y * z

forall x y : t, x - y == x + - y

forall x : t, x + - x == 0

forall x y : t, x * y == y * x

forall x y z : t, x * (y * z) == x * y * z

forall x y z : t, (x + y) * z == x * z + y * z

forall x y : t, x - y == x + - y

forall x : t, x + - x == 0

forall x y z : t, x * (y * z) == x * y * z

forall x y z : t, (x + y) * z == x * z + y * z

forall x y : t, x - y == x + - y

forall x : t, x + - x == 0

forall x y z : t, (x + y) * z == x * z + y * z

forall x y : t, x - y == x + - y

forall x : t, x + - x == 0

forall x y : t, x - y == x + - y

forall x : t, x + - x == 0
x, y:t

[|x + - y|] = [|x - y|]

forall x : t, x + - x == 0

forall x : t, x + - x == 0
exact add_opp_diag_r. Qed. Definition eqb x y := match ZnZ.compare x y with Eq => true | _ => false end.

forall x y : t, eqb x y = true <-> x == y

forall x y : t, eqb x y = true <-> x == y
x, y:t

eqb x y = true <-> x == y
x, y:t

match ZnZ.compare x y with | Eq => true | _ => false end = true <-> [|x|] = [|y|]
x, y:t

match [|x|] ?= [|y|] with | Eq => true | _ => false end = true <-> [|x|] = [|y|]
case Z.compare_spec; intuition; try discriminate. Qed.

forall x y : t, eqb x y = true -> x == y

forall x y : t, eqb x y = true -> x == y
now apply eqb_eq. Qed. End CyclicRing.