Library Coq.Numbers.NatInt.NZAdd


Some properties of the addition for modules implementing NZBasicFunsSig'

This file defines the NZAddProp functor type. This functor type is meant to be Included in a module implementing NZBasicFunsSig'.
This gives the following lemmas:
  • add_0_r, add_1_l, add_1_r
  • add_succ_r, add_succ_comm, add_comm
  • add_assoc
  • add_cancel_l, add_cancel_r
  • add_shuffle0, add_shuffle3 to rearrange sums of 3 elements
  • add_shuffle1, add_shuffle2 to rearrange sums of 4 elements
  • sub_1_r
From Coq.Numbers.NatInt Require Import NZAxioms NZBase.

Module Type NZAddProp (Import NZ : NZBasicFunsSig')(Import NZBase : NZBaseProp NZ).

Global Hint Rewrite
 pred_succ add_0_l add_succ_l mul_0_l mul_succ_l sub_0_r sub_succ_r : nz.
Global Hint Rewrite one_succ two_succ : nz'.
Ltac nzsimpl := autorewrite with nz.
Ltac nzsimpl' := autorewrite with nz nz'.

Theorem add_0_r : forall n, n + 0 == n.

Theorem add_succ_r : forall n m, n + S m == S (n + m).

Theorem add_succ_comm : forall n m, S n + m == n + S m.

Global Hint Rewrite add_0_r add_succ_r : nz.

Theorem add_comm : forall n m, n + m == m + n.

Theorem add_1_l : forall n, 1 + n == S n.

Theorem add_1_r : forall n, n + 1 == S n.

Global Hint Rewrite add_1_l add_1_r : nz.

Theorem add_assoc : forall n m p, n + (m + p) == (n + m) + p.

Theorem add_cancel_l : forall n m p, p + n == p + m <-> n == m.

Theorem add_cancel_r : forall n m p, n + p == m + p <-> n == m.

Theorem add_shuffle0 : forall n m p, n+m+p == n+p+m.

Theorem add_shuffle1 : forall n m p q, (n + m) + (p + q) == (n + p) + (m + q).

Theorem add_shuffle2 : forall n m p q, (n + m) + (p + q) == (n + q) + (m + p).

Theorem add_shuffle3 : forall n m p, n + (m + p) == m + (n + p).

Theorem sub_1_r : forall n, n - 1 == P n.

Global Hint Rewrite sub_1_r : nz.

End NZAddProp.