Lemmas about . Most of the rewriting support is ported from ss-reflect.
Symbols starting with vlib__ are internal.

Require Import Bool.
Require Import Arith.
Require Import ZArith.
Require Import String.
Require Import Vbase.

Set Implicit Arguments.

Comparison for nat

Fixpoint eqn_rec (x y: nat) {struct x} :=
   match x, y with
     | O, Otrue
     | S x, S yeqn_rec x y
     | _, _false
   end.

Definition eqn := match tt with tteqn_rec end.

Lemma eqnP: x y, reflect (x = y) (eqn x y).

Canonical Structure nat_eqMixin := EqMixin eqnP.
Canonical Structure nat_eqType := Eval hnf in EqType nat nat_eqMixin.

Lemma eqnE : eqn = (@eq_op _).

Three-valued logic


Three-valued Logic


Inductive bool3 := B3tt | B3ff | B3un.

Definition bool3_beq (x y : bool3) :=
  match x, y with
    | B3tt, B3tttrue
    | B3ff, B3fftrue
    | B3un, B3untrue
    | _, _false
  end.

Lemma bool3P : x y, reflect (x = y) (bool3_beq x y).

Definition negb3 (b:bool3) : bool3 :=
  match b with
   | B3ttB3ff
   | B3ffB3tt
   | B3unB3un
  end.

Definition andb3 (b1 b2:bool3) : bool3 :=
  match b1, b2 with
    | B3tt, _b2
    | B3ff, _B3ff
    | B3un, B3ffB3ff
    | B3un, _B3un
  end.

Definition orb3 (b1 b2:bool3) : bool3 :=
  match b1, b2 with
    | B3tt, _B3tt
    | B3ff, _b2
    | B3un, B3ttB3tt
    | B3un, _B3un
  end.

Definition bool2bool3 (b: bool) : bool3 :=
  if b then B3tt else B3ff.

Definition b3_moredef (b1 : bool3) (b2: bool) : bool :=
  match b1 with
    | B3ttb2
    | B3ffnegb b2
    | B3unfalse
  end.

Lemma andb3T: b, andb3 b B3tt = b.
Lemma andb3F: b, andb3 b B3ff = B3ff.
Lemma orb3T: b, orb3 b B3tt = B3tt.
Lemma orb3F: b, orb3 b B3ff = b.

Lemma negb3P: x, bool2bool3 (negb x) = negb3 (bool2bool3 x).

Lemma andb3P: x y, bool2bool3 (andb x y) = andb3 (bool2bool3 x) (bool2bool3 y).

Lemma orb3P: x y, bool2bool3 (orb x y) = orb3 (bool2bool3 x) (bool2bool3 y).

Lemma negb3_neg : x, negb3 (negb3 x) = x.

Lemma negb3_and : b1 b2, negb3 (andb3 b1 b2) = orb3 (negb3 b1) (negb3 b2).

Lemma negb3_or : b1 b2, negb3 (orb3 b1 b2) = andb3 (negb3 b1) (negb3 b2).

Hint Rewrite negb3_neg negb3_and negb3_or negb3P andb3P orb3P : vlib.

Natural numbers


Delimit Scope coq_nat_scope with coq_nat.

Notation "m <= n" := (le m n) : coq_nat_scope.
Notation "m < n" := (lt m n) : coq_nat_scope.
Notation "m >= n" := (ge m n) : coq_nat_scope.
Notation "m > n" := (gt m n) : coq_nat_scope.

Support for case analysis

Inductive LtSpec (A : Type) (le lt : AAProp) (x y: A): boolProp :=
  | LtSpecLt : lt x yLtSpec le lt x y true
  | LtSpecGe : le y xLtSpec le lt x y false.

Inductive CmpSpecFull (A : Type) (lt : AAProp) (x y: A)
  : boolboolboolboolboolboolcomparisonProp :=
  | CmpSpecLt : lt x yCmpSpecFull lt x y true true false false false false Lt
  | CmpSpecEq : x = yCmpSpecFull lt x y false true true true false true Eq
  | CmpSpecGt : lt y xCmpSpecFull lt x y false false false false true true Gt.

Boolean on nat
Definition len (m n:nat) := match tt with ttleb m n end.

Boolean < on nat.
Definition ltn m n := len (S m) n.


Overwrite Coq's standard notations

Notation "m <= n" := (len m n) : nat_scope.
Notation "m < n" := (ltn m n) : nat_scope.
Notation "m >= n" := (n m) (only parsing) : nat_scope.
Notation "m > n" := (n < m) (only parsing) : nat_scope.

Shorter names for properties of plus and minus.

Lemma add0n : n, 0 + n = n.
Lemma addSn : m n, S m + n = S (m + n).
Lemma add1n : n, 1 + n = S n.
Lemma addn0 : x, x + 0 = x.
Lemma addnS : m n, m + (S n) = S (m + n).
Lemma addn1 : n, n + 1 = S n.
Lemma addSnnS : m n, S m + n = m + (S n).

Lemma addnC : x y, x + y = y + x.
Lemma addnA : x y z, x + (y + z) = x + y + z.
Lemma addnCA : x y z, x + (y + z) = y + (x + z).
Lemma addnAC : x y z, x + y + z = x + z + y.

Lemma addn_eq0 : m n, (m + n == 0) = (m == 0) && (n == 0).

Lemma eqn_addl : p m n, (p + m == p + n) = (m == n).

Lemma eqn_addr : p m n, (m + p == n + p) = (m == n).

Lemma sub0n : x, 0 - x = 0.
Lemma subn0 : x, x - 0 = x.
Lemma subnn : x, x - x = 0.
Lemma subSS : n m, S m - S n = m - n.

Lemma subn_add2l : p m n, (p + m) - (p + n) = m - n.
Lemma subn_add2r : p m n, (m + p) - (n + p) = m - n.

Lemma subSnn : n, S n - n = 1.
Lemma subn_sub : m n p, (n - m) - p = n - (m + p).
Lemma subnAC : x y z, x - y - z = x - z - y.

Hint Rewrite
  add0n addSn addn0 addnS addn_eq0 eqn_addl eqn_addr
  sub0n subn0 subnn subSS subn_add2l subn_add2r : vlib.

Basic properties of comparisons


Lemma eq0S : n, (O == S n) = false.
Lemma eqS0 : n, (S n == O) = false.
Lemma eqSS : m n, (S m == S n) = (m == n).
Lemma eqnn : n : nat, (n == n).
Lemma eqnC : x y : nat, (x == y) = (y == x).

Hint Resolve eqnn.
Hint Rewrite eqnn eqSS eq0S eqS0 : vlib.

Lemma ltnE : m n, (m < n) = negb (n m).

Lemma lenE : m n, (m n) = negb (n < m).

Lemma lenE_sub : m n, (m n) = (m - n == 0).

Lemma ltnE_sub : m n, (m < n) = (m + 1 - n == 0).

Lemma leP : x y, LtSpec lt le x y (x y).

Lemma ltP : x y, LtSpec le lt x y (x < y).

Lemma ltn_correct: x y, (x < y)%coq_natx < y.

Lemma len_correct: x y, (x y)%coq_natx y.

Lemma ltn_complete: x y, x < y → (x < y)%coq_nat.

Lemma len_complete: x y, x y → (x y)%coq_nat.

Lemma lenP : x y, LtSpec ltn len x y (x y).

Lemma ltnP : x y, LtSpec len ltn x y (x < y).

Lemma le0n : n, 0 n.
Lemma len0 : n, (n 0) = (n == 0).
Lemma lenn : n, n n.
Lemma lenSn : n, n S n.

Lemma lt0Sn : n, 0 < S n.
Lemma ltn0 : n, (n < 0) = false.
Lemma ltnn : n, (n < n) = false.
Lemma ltnSn : n, (n < S n).

Lemma leSS : m n, (S m S n) = (m n).
Lemma ltSS : m n, (S m < S n) = (m < n).

Hint Resolve lenn lenSn ltnn ltnSn.

Lemma eq_len : m n, m = nm n.

Lemma ltnS : m n, (m < S n) = (m n).

Lemma leSn : m n, (S m n) = (m < n).

Hint Rewrite leSS le0n len0 lenn lenSn ltSS lt0Sn ltn0 ltnn ltnSn ltnS leSn : vlib.

Lemma lt0n : n, (0 < n) = negb (n == 0).
Lemma lt0n_neq0 : n, 0 < nnegb (n == 0).
Lemma eqn0_Nlt0n : n, n == 0 = negb (0 < n).
Lemma neq0_lt0n : n, n == 0 = false → 0 < n.
Hint Resolve lt0n_neq0 neq0_lt0n.

Lemma len_anti : m n, m nn mn = m.
Hint Immediate len_anti : lib.

Lemma eqn_leAle : m n, (m == n) = (m n) && (n m).

Lemma neqn_ltVlt : m n, negb (m == n) = (m < n) || (n < m).

Lemma len_eqVlt : m n, len m n = (m == n) || (m < n).

Lemma ltn_neqAle : m n, (m < n) = negb (m == n) && (m n).

Lemma len_trans : n m p, m nn pm p.

Lemma len_ltn_trans : n m p, m nn < pm < p.

Lemma ltn_len_trans : n m p, m < nn pm < p.

Lemma ltn_trans : n m p, m < nn < pm < p.

Lemma ltnW : m n, m < nm n.
Hint Resolve ltnW.

Lemma lenW : m n, m nm (S n).

Lemma len_total : m n, (m n) || (n m).

Lemma nat_comparenn: x, nat_compare x x = Eq.

Lemma cmpP : x y,
  CmpSpecFull ltn x y (x < y) (x y) (x == y) (y == x) (y < x) (y x) (nat_compare x y).

Monotonicity lemmata

Lemma len_add2l : p m n, (p + m p + n) = (m n).

Lemma ltn_add2l : p m n, (p + m < p + n) = (m < n).

Lemma len_add2r : p m n, (m + p n + p) = (m n).

Lemma ltn_add2r : p m n, (m + p < n + p) = (m < n).

Lemma len_add : m1 m2 n1 n2,
  m1 n1m2 n2m1 + m2 n1 + n2.

Lemma len_addr : m n, n n + m.

Lemma len_addl : m n, n m + n.

Lemma ltn_addr : m n p, m < nm < n + p.

Lemma ltn_addl : m n p, m < nm < p + n.

Lemma lt0_addn : m n, (0 < m + n) = (0 < m) || (0 < n).

Lemma lt0_subn : m n, (0 < n - m) = (m < n).

Lemma subn_eq0 : m n, (m - n == 0) = (m n).

Lemma len_sub_add : m n p, (m - n p) = (m n + p).

Lemma len_subr : m n, n - m n.

Lemma subnKC : m n, m nm + (n - m) = n.

Lemma subnK : m n, m n(n - m) + m = n.

Lemma addn_subA : m n p, p nm + (n - p) = m + n - p.

Lemma subn_subA : m n p, p nm - (n - p) = m + p - n.

Lemma subKn : m n, m nn - (n - m) = m.

Lemma len_subS : m n, m nS n - m = S (n - m).

Lemma ltn_subS : m n, m < nn - m = S (n - S m).

Lemma len_sub2r : p m n, (m n) → (m - p n - p).

Lemma len_sub2l : p m n, m np - n p - m.

Lemma len_sub2 : m1 m2 n1 n2,
  m1 m2n2 n1m1 - n1 m2 - n2.

Lemma ltn_sub2r : p m n, p < nm < nm - p < n - p.

Lemma ltn_sub2l : p m n, m < pm < np - n < p - m.

Lemma ltn_add_sub : m n p, (m + n < p) = (n < p - m).

Lemma lenE_diff: m n, m n k, n = m + k.

Lemma ltnE_diff: m n, m < n k, n = S (m + k).

Lemma len_subnE : m n, (m m - n) = (n == 0) || (m == 0).

Properties concerning multiplication


Lemma mul0n : n, 0 × n = 0.
Lemma muln0 : n, n × 0 = 0.
Lemma mul1n : n, 1 × n = n.
Lemma muln1 : n, n × 1 = n.

Lemma mulSn : m n, S m × n = n + m × n.
Lemma mulSnr : m n, S m × n = m × n + n.
Lemma mulnS : m n, m × S n = m + m × n.
Lemma mulnSr : m n, m × S n = m × n + m.

Lemma mulnC : m n, m × n = n × m.

Lemma muln_addl : x y z, (x + y) × z = x × z + y × z.

Lemma muln_addr : x y z, x × (y + z) = x × y + x × z.

Lemma muln_subl : x y z, (x - y) × z = x × z - y × z.

Lemma muln_subr : x y z, x × (y - z) = x × y - x × z.

Lemma mulnA : x y z, x × (y × z) = x × y × z.

Lemma mulnCA : x y z, x × (y × z) = y × (x × z).

Lemma mulnAC : x y z, x × y × z = x × z × y.

Lemma eqn_mul0 : m n, (m × n == 0) = (m == 0) || (n == 0).

Lemma eqn_mul1 : m n, (m × n == 1) = (m == 1) && (n == 1).

Lemma lt0_muln : m n, (0 < m × n) = (0 < m) && (0 < n).

Lemma len_pmull : m n, 0 < nm n × m.

Lemma len_pmulr : m n, 0 < nm m × n.

Lemma len_mul2l : m n1 n2, (m × n1 m × n2) = (m == 0) || (n1 n2).

Lemma len_mul2r : m n1 n2, (n1 × m n2 × m) = (m == 0) || (n1 n2).

Lemma len_mul : m1 m2 n1 n2, m1 n1m2 n2m1 × m2 n1 × n2.

Lemma eqn_mul2l : m n1 n2, (m × n1 == m × n2) = (m == 0) || (n1 == n2).

Lemma eqn_mul2r : m n1 n2, (n1 × m == n2 × m) = (m == 0) || (n1 == n2).

Lemma len_pmul2l : m n1 n2, 0 < m(m × n1 m × n2) = (n1 n2).
Implicit Arguments len_pmul2l [m n1 n2].

Lemma len_pmul2r : m n1 n2, 0 < m(n1 × m n2 × m) = (n1 n2).
Implicit Arguments len_pmul2r [m n1 n2].

Lemma eqn_pmul2l : m n1 n2, 0 < m(m × n1 == m × n2) = (n1 == n2).
Implicit Arguments eqn_pmul2l [m n1 n2].

Lemma eqn_pmul2r : m n1 n2, 0 < m(n1 × m == n2 × m) = (n1 == n2).
Implicit Arguments eqn_pmul2r [m n1 n2].

Lemma ltn_mul2l : m n1 n2, (m × n1 < m × n2) = (0 < m) && (n1 < n2).

Lemma ltn_mul2r : m n1 n2, (n1 × m < n2 × m) = (0 < m) && (n1 < n2).

Lemma ltn_pmul2l : m n1 n2, 0 < m(m × n1 < m × n2) = (n1 < n2).
Implicit Arguments ltn_pmul2l [m n1 n2].

Lemma ltn_pmul2r : m n1 n2, 0 < m(n1 × m < n2 × m) = (n1 < n2).
Implicit Arguments ltn_pmul2r [m n1 n2].

Lemma ltn_mul : m1 m2 n1 n2, m1 < n1m2 < n2m1 × m2 < n1 × n2.

Lemma ltn_Pmull : m n, 1 < n → 0 < mm < n × m.

Lemma ltn_Pmulr : m n, 1 < n → 0 < mm < m × n.

auto setup

Lemma lenS : m n, n mn S m.

Lemma ltn_of_leS : n m : nat, (n < m) → (S n m).

Lemma leSn_of_lt : n m : nat, (S n m) → (n < m).

Lemma len_mul2l_imp : m n1 n2, n1 n2 → (m × n1 m × n2).

Lemma len_mul2r_imp : m n1 n2, n1 n2 → (n1 × m n2 × m).

Hint Resolve lenS len_addl len_addr len_trans len_add2r len_add2l len_sub2r len_sub2l
             len_mul2l_imp len_mul2r_imp : vlib.
Hint Immediate ltn_of_leS leSn_of_lt: vlib.

Hint Rewrite add0n addSn addn0 addnS addn_eq0 eqn_addl eqn_addr
  sub0n subn0 subnn subSS subn_add2l subn_add2r
  nat_comparenn
  len_add2l ltn_add2l len_add2r ltn_add2r len_addr len_addl
  lt0_addn lt0_subn len_sub_add len_subnE
  mul0n muln0 mul1n muln1

  muln_addl muln_addr muln_subl muln_subr
  eqn_mul0 eqn_mul1 lt0_muln
  len_mul2l len_mul2r eqn_mul2l eqn_mul2r ltn_mul2l ltn_mul2r : vlib.

Hint Rewrite addnA mulnA : vlibA.

iter f n applies n times the function f. This definition is not-tail-recursive: it is more convenient for proofs.
Fixpoint iter (A : Type) n (f : AA) (v : A): A :=
  match n with
    | Ov
    | S nf (iter n f v)
  end.

Tail-recursive version

Fixpoint iter_rec (A : Type) n (f : AA) (v : A): A :=
  match n with
    | Ov
    | S niter_rec n f (f v)
  end.

Lemma iter_recS: A n f (v: A), iter_rec (S n) f v = f (iter_rec n f v).

Lemma iter_recE: A n f (v: A), iter_rec n f v = iter n f v.

Integers


Section ZZZ.

Local Open Scope Z_scope.

Lemma Zplus_mod_same: a n : Z, (a + n) mod n = a mod n.

Lemma Zmod_too_small:
   a n, -n a < 0 → a mod n = n + a.

Lemma Zmod_too_big:
   a n, n a < n + na mod n = a - n.

Lemma Zopp_mod_idemp: a m, - (a mod m) mod m = - a mod m.

Lemma Zmod_opp_expand: a b, - a mod b = (b - a) mod b.

Lemma Zmod_eqD :
   a b m (H: a mod m = b mod m) (NEQ: m 0), k, a = b + k × m.

Lemma ZdoubleE: x, Zdouble x = x × 2.

Lemma ZltP:
   x y, LtSpec Zle Zlt x y (Zlt_bool x y).

Lemma ZleP:
   x y, LtSpec Zlt Zle x y (Zle_bool x y).

Lemma ZcmpP:
   x y, CmpSpecFull Zlt x y
    (Zlt_bool x y) (Zle_bool x y) (Zeq_bool x y) (Zeq_bool y x)
    (Zlt_bool y x) (Zle_bool y x) (Zcompare x y).

End ZZZ.

Hint Rewrite
  Zmod_0_l Z_mod_same_full
  Zplus_mod_idemp_l Zplus_mod_idemp_r
  Zminus_mod_idemp_l Zminus_mod_idemp_r Zopp_mod_idemp : vlib.

This page has been generated by coqdoc