Skip to content
Snippets Groups Projects
numbers.v 56 KiB
Newer Older
(** This file collects some trivial facts on the Coq types [nat] and [N] for
natural numbers, and the type [Z] for integers. It also declares some useful
notations. *)
From Coq Require Export EqdepFacts PArith NArith ZArith NPeano.
From Coq Require Import QArith Qcanon.
From stdpp Require Export base decidable option sprop.
Ralf Jung's avatar
Ralf Jung committed
From stdpp Require Import options.
Local Open Scope nat_scope.
Robbert Krebbers's avatar
Robbert Krebbers committed

Global Instance comparison_eq_dec : EqDecision comparison.
Proof. solve_decision. Defined.
(** * Notations and properties of [nat] *)
Global Arguments minus !_ !_ / : assert.
Global Arguments Nat.max : simpl nomatch.
Reserved Notation "x ≤ y ≤ z" (at level 70, y at next level).
Reserved Notation "x ≤ y < z" (at level 70, y at next level).
Reserved Notation "x < y < z" (at level 70, y at next level).
Reserved Notation "x < y ≤ z" (at level 70, y at next level).
Reserved Notation "x ≤ y ≤ z ≤ z'"
  (at level 70, y at next level, z at next level).
Infix "≤" := le : nat_scope.
Notation "x ≤ y ≤ z" := (x  y  y  z)%nat : nat_scope.
Notation "x ≤ y < z" := (x  y  y < z)%nat : nat_scope.
Notation "x < y ≤ z" := (x < y  y  z)%nat : nat_scope.
Notation "x ≤ y ≤ z ≤ z'" := (x  y  y  z  z  z')%nat : nat_scope.
Notation "(≤)" := le (only parsing) : nat_scope.
Notation "(<)" := lt (only parsing) : nat_scope.

Infix "`div`" := Nat.div (at level 35) : nat_scope.
Infix "`mod`" := Nat.modulo (at level 35) : nat_scope.
Infix "`max`" := Nat.max (at level 35) : nat_scope.
Infix "`min`" := Nat.min (at level 35) : nat_scope.
Global Instance nat_eq_dec: EqDecision nat := eq_nat_dec.
Global Instance nat_le_dec: RelDecision le := le_dec.
Global Instance nat_lt_dec: RelDecision lt := lt_dec.
Global Instance nat_inhabited: Inhabited nat := populate 0%nat.
Global Instance S_inj: Inj (=) (=) S.
Proof. by injection 1. Qed.
Global Instance nat_le_po: PartialOrder ().
Proof. repeat split; repeat intro; auto with lia. Qed.
Global Instance nat_le_total: Total ().
Proof. repeat intro; lia. Qed.
Global Instance nat_le_pi:  x y : nat, ProofIrrel (x  y).
Proof.
  assert ( x y (p : x  y) y' (q : x  y'),
    y = y'  eq_dep nat (le x) y p y' q) as aux.
  { fix FIX 3. intros x ? [|y p] ? [|y' q].
    - clear FIX. intros; exfalso; auto with lia.
    - clear FIX. intros; exfalso; auto with lia.
    - injection 1. intros Hy. by case (FIX x y p y' q Hy). }
  by apply (Eqdep_dec.eq_dep_eq_dec (λ x y, decide (x = y))), aux.
Global Instance nat_lt_pi:  x y : nat, ProofIrrel (x < y).
Proof. unfold lt. apply _. Qed.
Robbert Krebbers's avatar
Robbert Krebbers committed
Lemma nat_le_sum (x y : nat) : x  y   z, y = x + z.
Proof. split; [exists (y - x); lia | intros [z ->]; lia]. Qed.
Lemma Nat_lt_succ_succ n : n < S (S n).
Proof. auto with arith. Qed.
Lemma Nat_mul_split_l n x1 x2 y1 y2 :
  x2 < n  y2 < n  x1 * n + x2 = y1 * n + y2  x1 = y1  x2 = y2.
Proof.
  intros Hx2 Hy2 E. cut (x1 = y1); [intros; subst;lia |].
  revert y1 E. induction x1; simpl; intros [|?]; simpl; auto with lia.
Qed.
Lemma Nat_mul_split_r n x1 x2 y1 y2 :
  x1 < n  y1 < n  x1 + x2 * n = y1 + y2 * n  x1 = y1  x2 = y2.
Proof. intros. destruct (Nat_mul_split_l n x2 x1 y2 y1); auto with lia. Qed.
Notation lcm := Nat.lcm.
Notation divide := Nat.divide.
Notation "( x | y )" := (divide x y) : nat_scope.
Global Instance Nat_divide_dec : RelDecision Nat.divide.
  refine (λ x y, cast_if (decide (lcm x y = y))); by rewrite Nat.divide_lcm_iff.
Global Instance: PartialOrder divide.
Proof.
  repeat split; try apply _. intros ??. apply Nat.divide_antisym_nonneg; lia.
Qed.
Global Hint Extern 0 (_ | _) => reflexivity : core.
Lemma Nat_divide_ne_0 x y : (x | y)  y  0  x  0.
Proof. intros Hxy Hy ->. by apply Hy, Nat.divide_0_l. Qed.

Lemma Nat_iter_S {A} n (f: A  A) x : Nat.iter (S n) f x = f (Nat.iter n f x).
Proof. done. Qed.
Lemma Nat_iter_S_r {A} n (f: A  A) x : Nat.iter (S n) f x = Nat.iter n f (f x).
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. induction n; by f_equal/=. Qed.
Lemma Nat_iter_add {A} n1 n2 (f : A  A) x :
  Nat.iter (n1 + n2) f x = Nat.iter n1 f (Nat.iter n2 f x).
Proof. induction n1; by f_equal/=. Qed.
Ralf Jung's avatar
Ralf Jung committed
Lemma Nat_iter_mul {A} n1 n2 (f : A  A) x :
  Nat.iter (n1 * n2) f x = Nat.iter n1 (Nat.iter n2 f) x.
Proof.
  intros. induction n1 as [|n1 IHn1]; [done|].
  simpl. by rewrite Nat_iter_add, IHn1.
Qed.
Robbert Krebbers's avatar
Robbert Krebbers committed
Lemma Nat_iter_ind {A} (P : A  Prop) f x k :
Robbert Krebbers's avatar
Robbert Krebbers committed
  P x  ( y, P y  P (f y))  P (Nat.iter k f x).
Proof. induction k; simpl; auto. Qed.
(** * Notations and properties of [positive] *)
Local Open Scope positive_scope.
Typeclasses Opaque Pos.le.
Typeclasses Opaque Pos.lt.

Infix "≤" := Pos.le : positive_scope.
Notation "x ≤ y ≤ z" := (x  y  y  z) : positive_scope.
Notation "x ≤ y < z" := (x  y  y < z) : positive_scope.
Notation "x < y ≤ z" := (x < y  y  z) : positive_scope.
Notation "x ≤ y ≤ z ≤ z'" := (x  y  y  z  z  z') : positive_scope.
Notation "(≤)" := Pos.le (only parsing) : positive_scope.
Notation "(<)" := Pos.lt (only parsing) : positive_scope.
Robbert Krebbers's avatar
Robbert Krebbers committed
Notation "(~0)" := xO (only parsing) : positive_scope.
Notation "(~1)" := xI (only parsing) : positive_scope.

Global Arguments Pos.of_nat : simpl never.
Global Arguments Pmult : simpl never.
Global Instance positive_eq_dec: EqDecision positive := Pos.eq_dec.
Global Instance positive_le_dec: RelDecision Pos.le.
Proof. refine (λ x y, decide ((x ?= y)  Gt)). Defined.
Global Instance positive_lt_dec: RelDecision Pos.lt.
Proof. refine (λ x y, decide ((x ?= y) = Lt)). Defined.
Global Instance positive_le_total: Total Pos.le.
Proof. repeat intro; lia. Qed.

Global Instance positive_inhabited: Inhabited positive := populate 1.
Global Instance maybe_xO : Maybe xO := λ p, match p with p~0 => Some p | _ => None end.
Global Instance maybe_xI : Maybe xI := λ p, match p with p~1 => Some p | _ => None end.
Global Instance xO_inj : Inj (=) (=) (~0).
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. by injection 1. Qed.
Global Instance xI_inj : Inj (=) (=) (~1).
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. by injection 1. Qed.

(** Since [positive] represents lists of bits, we define list operations
on it. These operations are in reverse, as positives are treated as snoc
lists instead of cons lists. *)
Fixpoint Papp (p1 p2 : positive) : positive :=
  match p2 with
  | 1 => p1
  | p2~0 => (Papp p1 p2)~0
  | p2~1 => (Papp p1 p2)~1
  end.
Infix "++" := Papp : positive_scope.
Notation "(++)" := Papp (only parsing) : positive_scope.
Notation "( p ++.)" := (Papp p) (only parsing) : positive_scope.
Notation "(.++ q )" := (λ p, Papp p q) (only parsing) : positive_scope.

Fixpoint Preverse_go (p1 p2 : positive) : positive :=
  match p2 with
  | 1 => p1
  | p2~0 => Preverse_go (p1~0) p2
  | p2~1 => Preverse_go (p1~1) p2
  end.
Definition Preverse : positive  positive := Preverse_go 1.

Robbert Krebbers's avatar
Robbert Krebbers committed
Global Instance Papp_1_l : LeftId (=) 1 (++).
Proof. intros p. by induction p; intros; f_equal/=. Qed.
Robbert Krebbers's avatar
Robbert Krebbers committed
Global Instance Papp_1_r : RightId (=) 1 (++).
Robbert Krebbers's avatar
Robbert Krebbers committed
Global Instance Papp_assoc : Assoc (=) (++).
Proof. intros ?? p. by induction p; intros; f_equal/=. Qed.
Global Instance Papp_inj p : Inj (=) (=) (.++ p).
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. intros ???. induction p; simplify_eq; auto. Qed.

Lemma Preverse_go_app p1 p2 p3 :
  Preverse_go p1 (p2 ++ p3) = Preverse_go p1 p3 ++ Preverse_go 1 p2.
Proof.
  revert p3 p1 p2.
  cut ( p1 p2 p3, Preverse_go (p2 ++ p3) p1 = p2 ++ Preverse_go p3 p1).
  { by intros go p3; induction p3; intros p1 p2; simpl; auto; rewrite <-?go. }
  intros p1; induction p1 as [p1 IH|p1 IH|]; intros p2 p3; simpl; auto.
  - apply (IH _ (_~1)).
  - apply (IH _ (_~0)).
Lemma Preverse_app p1 p2 : Preverse (p1 ++ p2) = Preverse p2 ++ Preverse p1.
Proof. unfold Preverse. by rewrite Preverse_go_app. Qed.
Lemma Preverse_xO p : Preverse (p~0) = (1~0) ++ Preverse p.
Proof Preverse_app p (1~0).
Lemma Preverse_xI p : Preverse (p~1) = (1~1) ++ Preverse p.
Proof Preverse_app p (1~1).

Lemma Preverse_involutive p :
  Preverse (Preverse p) = p.
Proof.
  induction p as [p IH|p IH|]; simpl.
  - by rewrite Preverse_xI, Preverse_app, IH.
  - by rewrite Preverse_xO, Preverse_app, IH.
  - reflexivity.
Qed.
Global Instance Preverse_inj : Inj (=) (=) Preverse.
Proof.
  intros p q eq.
  rewrite <- (Preverse_involutive p).
  rewrite <- (Preverse_involutive q).
  by rewrite eq.
Qed.

Fixpoint Plength (p : positive) : nat :=
  match p with 1 => 0%nat | p~0 | p~1 => S (Plength p) end.
Lemma Papp_length p1 p2 : Plength (p1 ++ p2) = (Plength p2 + Plength p1)%nat.
Proof. by induction p2; f_equal/=. Qed.
Robbert Krebbers's avatar
Robbert Krebbers committed
Lemma Plt_sum (x y : positive) : x < y   z, y = x + z.
Proof.
  split.
  - exists (y - x)%positive. symmetry. apply Pplus_minus. lia.
  - intros [z ->]. lia.
Qed.

(** Duplicate the bits of a positive, i.e. 1~0~1 -> 1~0~0~1~1 and
    1~1~0~0 -> 1~1~1~0~0~0~0 *)
Fixpoint Pdup (p : positive) : positive :=
  match p with
  | 1 => 1
  | p'~0 => (Pdup p')~0~0
  | p'~1 => (Pdup p')~1~1
  end.

Lemma Pdup_app p q :
  Pdup (p ++ q) = Pdup p ++ Pdup q.
Proof.
  revert p.
  induction q as [p IH|p IH|]; intros q; simpl.
  - by rewrite IH.
  - by rewrite IH.
  - reflexivity.
Qed.

Lemma Pdup_suffix_eq p q s1 s2 :
  s1~1~0 ++ Pdup p = s2~1~0 ++ Pdup q  p = q.
Proof.
  revert q.
  induction p as [p IH|p IH|]; intros [q|q|] eq; simplify_eq/=.
  - by rewrite (IH q).
  - by rewrite (IH q).
  - reflexivity.
Qed.

Global Instance Pdup_inj : Inj (=) (=) Pdup.
Proof.
  intros p q eq.
  apply (Pdup_suffix_eq _ _ 1 1).
  by rewrite eq.
Qed.

Lemma Preverse_Pdup p :
  Preverse (Pdup p) = Pdup (Preverse p).
Proof.
  induction p as [p IH|p IH|]; simpl.
  - rewrite 3!Preverse_xI.
    rewrite (assoc_L (++)).
    rewrite IH.
    rewrite Pdup_app.
    reflexivity.
  - rewrite 3!Preverse_xO.
    rewrite (assoc_L (++)).
    rewrite IH.
    rewrite Pdup_app.
    reflexivity.
  - reflexivity.
Qed.

Local Close Scope positive_scope.

(** * Notations and properties of [N] *)
Typeclasses Opaque N.le.
Typeclasses Opaque N.lt.

Robbert Krebbers's avatar
Robbert Krebbers committed
Infix "≤" := N.le : N_scope.
Notation "x ≤ y ≤ z" := (x  y  y  z)%N : N_scope.
Notation "x ≤ y < z" := (x  y  y < z)%N : N_scope.
Notation "x < y ≤ z" := (x < y  y  z)%N : N_scope.
Notation "x ≤ y ≤ z ≤ z'" := (x  y  y  z  z  z')%N : N_scope.
Robbert Krebbers's avatar
Robbert Krebbers committed
Notation "(≤)" := N.le (only parsing) : N_scope.
Notation "(<)" := N.lt (only parsing) : N_scope.
Infix "`div`" := N.div (at level 35) : N_scope.
Infix "`mod`" := N.modulo (at level 35) : N_scope.
Infix "`max`" := N.max (at level 35) : N_scope.
Infix "`min`" := N.min (at level 35) : N_scope.
Global Arguments N.add : simpl never.
Global Instance Npos_inj : Inj (=) (=) Npos.
Robbert Krebbers's avatar
Robbert Krebbers committed
Proof. by injection 1. Qed.

Global Instance N_eq_dec: EqDecision N := N.eq_dec.
Global Program Instance N_le_dec : RelDecision N.le := λ x y,
  match N.compare x y with Gt => right _ | _ => left _ end.
Solve Obligations with naive_solver.
Global Program Instance N_lt_dec : RelDecision N.lt := λ x y,
  match N.compare x y with Lt => left _ | _ => right _ end.
Solve Obligations with naive_solver.
Global Instance N_inhabited: Inhabited N := populate 1%N.
Global Instance N_lt_pi x y : ProofIrrel (x < y)%N.
Proof. unfold N.lt. apply _. Qed.

Global Instance N_le_po: PartialOrder ()%N.
  repeat split; red; [apply N.le_refl | apply N.le_trans | apply N.le_antisymm].
Global Instance N_le_total: Total ()%N.
Proof. repeat intro; lia. Qed.

Global Hint Extern 0 (_  _)%N => reflexivity : core.
Robbert Krebbers's avatar
Robbert Krebbers committed

(** * Notations and properties of [Z] *)
Local Open Scope Z_scope.
Typeclasses Opaque Z.le.
Typeclasses Opaque Z.lt.

Robbert Krebbers's avatar
Robbert Krebbers committed
Infix "≤" := Z.le : Z_scope.
Notation "x ≤ y ≤ z" := (x  y  y  z) : Z_scope.
Notation "x ≤ y < z" := (x  y  y < z) : Z_scope.
Notation "x < y < z" := (x < y  y < z) : Z_scope.
Notation "x < y ≤ z" := (x < y  y  z) : Z_scope.
Notation "x ≤ y ≤ z ≤ z'" := (x  y  y  z  z  z') : Z_scope.
Robbert Krebbers's avatar
Robbert Krebbers committed
Notation "(≤)" := Z.le (only parsing) : Z_scope.
Notation "(<)" := Z.lt (only parsing) : Z_scope.
Infix "`div`" := Z.div (at level 35) : Z_scope.
Infix "`mod`" := Z.modulo (at level 35) : Z_scope.
Infix "`quot`" := Z.quot (at level 35) : Z_scope.
Infix "`rem`" := Z.rem (at level 35) : Z_scope.
Infix "≪" := Z.shiftl (at level 35) : Z_scope.
Infix "≫" := Z.shiftr (at level 35) : Z_scope.
Infix "`max`" := Z.max (at level 35) : Z_scope.
Infix "`min`" := Z.min (at level 35) : Z_scope.
Global Instance Zpos_inj : Inj (=) (=) Zpos.
Global Instance Zneg_inj : Inj (=) (=) Zneg.
Global Instance Z_of_nat_inj : Inj (=) (=) Z.of_nat.
Proof. intros n1 n2. apply Nat2Z.inj. Qed.

Global Instance Z_eq_dec: EqDecision Z := Z.eq_dec.
Global Instance Z_le_dec: RelDecision Z.le := Z_le_dec.
Global Instance Z_lt_dec: RelDecision Z.lt := Z_lt_dec.
Global Instance Z_ge_dec: RelDecision Z.ge := Z_ge_dec.
Global Instance Z_gt_dec: RelDecision Z.gt := Z_gt_dec.
Global Instance Z_inhabited: Inhabited Z := populate 1.
Global Instance Z_lt_pi x y : ProofIrrel (x < y).
Proof. unfold Z.lt. apply _. Qed.

Global Instance Z_le_po : PartialOrder ().
  repeat split; red; [apply Z.le_refl | apply Z.le_trans | apply Z.le_antisymm].
Global Instance Z_le_total: Total Z.le.
Proof. repeat intro; lia. Qed.

Lemma Z_pow_pred_r n m : 0 < m  n * n ^ (Z.pred m) = n ^ m.
Proof.
  intros. rewrite <-Z.pow_succ_r, Z.succ_pred; [done|]. by apply Z.lt_le_pred.
Qed.
Lemma Z_quot_range_nonneg k x y : 0  x < k  0 < y  0  x `quot` y < k.
Proof.
  intros [??] ?.
  destruct (decide (y = 1)); subst; [rewrite Z.quot_1_r; auto |].
  destruct (decide (x = 0)); subst; [rewrite Z.quot_0_l; auto with lia |].
  split; [apply Z.quot_pos; lia|].
  trans x; auto. apply Z.quot_lt; lia.
Robbert Krebbers's avatar
Robbert Krebbers committed

Global Arguments Z.pred : simpl never.
Global Arguments Z.succ : simpl never.
Global Arguments Z.of_nat : simpl never.
Global Arguments Z.to_nat : simpl never.
Global Arguments Z.mul : simpl never.
Global Arguments Z.add : simpl never.
Global Arguments Z.sub : simpl never.
Global Arguments Z.opp : simpl never.
Global Arguments Z.pow : simpl never.
Global Arguments Z.div : simpl never.
Global Arguments Z.modulo : simpl never.
Global Arguments Z.quot : simpl never.
Global Arguments Z.rem : simpl never.
Global Arguments Z.shiftl : simpl never.
Global Arguments Z.shiftr : simpl never.
Global Arguments Z.gcd : simpl never.
Global Arguments Z.lcm : simpl never.
Global Arguments Z.min : simpl never.
Global Arguments Z.max : simpl never.
Global Arguments Z.lor : simpl never.
Global Arguments Z.land : simpl never.
Global Arguments Z.lxor : simpl never.
Global Arguments Z.lnot : simpl never.
Global Arguments Z.square : simpl never.
Global Arguments Z.abs : simpl never.
Lemma Z_to_nat_neq_0_pos x : Z.to_nat x  0%nat  0 < x.
Proof. by destruct x. Qed.
Lemma Z_to_nat_neq_0_nonneg x : Z.to_nat x  0%nat  0  x.
Proof. by destruct x. Qed.
Lemma Z_mod_pos x y : 0 < y  0  x `mod` y.
Global Hint Resolve Z.lt_le_incl : zpos.
Global Hint Resolve Z.add_nonneg_pos Z.add_pos_nonneg Z.add_nonneg_nonneg : zpos.
Global Hint Resolve Z.mul_nonneg_nonneg Z.mul_pos_pos : zpos.
Global Hint Resolve Z.pow_pos_nonneg Z.pow_nonneg: zpos.
Global Hint Resolve Z_mod_pos Z.div_pos : zpos.
Global Hint Extern 1000 => lia : zpos.
Robbert Krebbers's avatar
Robbert Krebbers committed
Lemma Z_to_nat_nonpos x : x  0  Z.to_nat x = 0%nat.
Proof. destruct x; simpl; auto using Z2Nat.inj_neg. by intros []. Qed.
Lemma Z2Nat_inj_pow (x y : nat) : Z.of_nat (x ^ y) = (Z.of_nat x) ^ (Z.of_nat y).
Robbert Krebbers's avatar
Robbert Krebbers committed
  induction y as [|y IH]; [by rewrite Z.pow_0_r, Nat.pow_0_r|].
  by rewrite Nat.pow_succ_r, Nat2Z.inj_succ, Z.pow_succ_r,
    Nat2Z.inj_mul, IH by auto with zpos.
Robbert Krebbers's avatar
Robbert Krebbers committed
Lemma Nat2Z_divide n m : (Z.of_nat n | Z.of_nat m)  (n | m)%nat.
Proof.
  split.
  - rewrite <-(Nat2Z.id m) at 2; intros [i ->]; exists (Z.to_nat i).
Robbert Krebbers's avatar
Robbert Krebbers committed
    destruct (decide (0  i)%Z).
    { by rewrite Z2Nat.inj_mul, Nat2Z.id by lia. }
    by rewrite !Z_to_nat_nonpos by auto using Z.mul_nonpos_nonneg with lia.
  - intros [i ->]. exists (Z.of_nat i). by rewrite Nat2Z.inj_mul.
Robbert Krebbers's avatar
Robbert Krebbers committed
Qed.
Lemma Z2Nat_divide n m :
  0  n  0  m  (Z.to_nat n | Z.to_nat m)%nat  (n | m).
Proof. intros. by rewrite <-Nat2Z_divide, !Z2Nat.id by done. Qed.
Lemma Nat2Z_inj_div x y : Z.of_nat (x `div` y) = (Z.of_nat x) `div` (Z.of_nat y).
Proof.
  destruct (decide (y = 0%nat)); [by subst; destruct x |].
  apply Z.div_unique with (Z.of_nat $ x `mod` y)%nat.
  { left. rewrite <-(Nat2Z.inj_le 0), <-Nat2Z.inj_lt.
    apply Nat.mod_bound_pos; lia. }
  by rewrite <-Nat2Z.inj_mul, <-Nat2Z.inj_add, <-Nat.div_mod.
Qed.
Lemma Nat2Z_inj_mod x y : Z.of_nat (x `mod` y) = (Z.of_nat x) `mod` (Z.of_nat y).
Proof.
  destruct (decide (y = 0%nat)); [by subst; destruct x |].
  apply Z.mod_unique with (Z.of_nat $ x `div` y)%nat.
Loading
Loading full blame...