Library Faux
Require Export Min.
Require Export Arith.
Require Export Reals.
Require Export Zpower.
Require Export ZArith.
Require Export Zcomplements.
Require Export sTactic.
Hint Resolve R1_neq_R0: real.
Theorem minus_minus : forall a b : nat, a <= b -> b - (b - a) = a.
Theorem lte_comp_mult :
forall p q r t : nat, p <= q -> r <= t -> p * r <= q * t.
Hint Resolve lte_comp_mult: arith.
Theorem le_refl_eq : forall n m : nat, n = m -> n <= m.
Theorem lt_le_pred : forall n m : nat, n < m -> n <= pred m.
Theorem lt_comp_mult_l : forall p q r : nat, 0 < p -> q < r -> p * q < p * r.
Hint Resolve lt_comp_mult_l: arith.
Theorem lt_comp_mult_r : forall p q r : nat, 0 < p -> q < r -> q * p < r * p.
Hint Resolve lt_comp_mult_r: arith.
Theorem lt_comp_mult : forall p q r s : nat, p < q -> r < s -> p * r < q * s.
Hint Resolve lt_comp_mult: arith.
Theorem mult_eq_inv : forall n m p : nat, 0 < n -> n * m = n * p -> m = p.
Definition natEq : forall n m : nat, {n = m} + {n <> m}.
Theorem notEqLt : forall n : nat, 0 < n -> n <> 0.
Hint Resolve notEqLt: arith.
Theorem lt_next : forall n m : nat, n < m -> m = S n \/ S n < m.
Theorem le_next : forall n m : nat, n <= m -> m = n \/ S n <= m.
Theorem min_or :
forall n m : nat, min n m = n /\ n <= m \/ min n m = m /\ m < n.
Theorem minus_inv_lt_aux : forall n m : nat, n - m = 0 -> n - S m = 0.
Theorem minus_inv_lt : forall n m : nat, m <= n -> m - n = 0.
Theorem minus_le : forall m n p q : nat, m <= n -> p <= q -> m - q <= n - p.
Theorem lt_minus_inv : forall n m p : nat, n <= p -> m < n -> p - n < p - m.
Theorem lt_mult_anti_compatibility :
forall n n1 n2 : nat, 0 < n -> n * n1 < n * n2 -> n1 < n2.
Theorem le_mult_anti_compatibility :
forall n n1 n2 : nat, 0 < n -> n * n1 <= n * n2 -> n1 <= n2.
Theorem min_n_0 : forall n : nat, min n 0 = 0.
Hint Resolve Rabs_pos: real.
Theorem Rlt_Rminus_ZERO : forall r1 r2 : R, (r2 < r1)%R -> (0 < r1 - r2)%R.
Hint Resolve Rlt_Rminus_ZERO: real.
Theorem Rabsolu_left1 : forall a : R, (a <= 0)%R -> Rabs a = (- a)%R.
Theorem RmaxLess1 : forall r1 r2 : R, (r1 <= Rmax r1 r2)%R.
Theorem RmaxLess2 : forall r1 r2 : R, (r2 <= Rmax r1 r2)%R.
Theorem RmaxSym : forall p q : R, Rmax p q = Rmax q p.
Theorem RmaxAbs :
forall p q r : R,
(p <= q)%R -> (q <= r)%R -> (Rabs q <= Rmax (Rabs p) (Rabs r))%R.
Theorem Rabsolu_Zabs : forall z : Z, Rabs (IZR z) = IZR (Zabs z).
Theorem RmaxRmult :
forall p q r : R, (0 <= r)%R -> Rmax (r * p) (r * q) = (r * Rmax p q)%R.
Theorem Rle_R0_Ropp : forall p : R, (p <= 0)%R -> (0 <= - p)%R.
Theorem Rlt_R0_Ropp : forall p : R, (p < 0)%R -> (0 < - p)%R.
Hint Resolve Rle_R0_Ropp Rlt_R0_Ropp: real.
Theorem convert_not_O : forall p : positive, nat_of_P p <> 0.
Hint Resolve convert_not_O: zarith arith.
Hint Resolve Zlt_le_weak Zle_not_gt Zgt_irrefl Zlt_irrefl Zle_not_lt
Zlt_not_le Zlt_asym inj_lt inj_le: zarith.
Theorem inj_abs :
forall x : Z, (0 <= x)%Z -> Z_of_nat (Zabs_nat x) = x.
Theorem inject_nat_convert :
forall (p : Z) (q : positive),
p = Zpos q -> Z_of_nat (nat_of_P q) = p.
Hint Resolve inj_le inj_lt: zarith.
Theorem ZleLe : forall x y : nat, (Z_of_nat x <= Z_of_nat y)%Z -> x <= y.
Theorem inject_nat_eq : forall x y : nat, Z_of_nat x = Z_of_nat y -> x = y.
Theorem Zcompare_EGAL :
forall p q : Z, (p ?= q)%Z = Datatypes.Eq -> p = q.
Theorem Zlt_Zopp : forall x y : Z, (x < y)%Z -> (- y < - x)%Z.
Hint Resolve Zlt_Zopp: zarith.
Theorem Zle_Zopp : forall x y : Z, (x <= y)%Z -> (- y <= - x)%Z.
Hint Resolve Zle_Zopp: zarith.
Theorem absolu_INR : forall n : nat, Zabs_nat (Z_of_nat n) = n.
Theorem absolu_Zopp : forall p : Z, Zabs_nat (- p) = Zabs_nat p.
Theorem Zabs_absolu : forall z : Z, Zabs z = Z_of_nat (Zabs_nat z).
Theorem absolu_comp_mult :
forall p q : Z, Zabs_nat (p * q) = Zabs_nat p * Zabs_nat q.
Theorem Zmin_sym : forall m n : Z, Zmin n m = Zmin m n.
Theorem Zpower_nat_O : forall z : Z, Zpower_nat z 0 = Z_of_nat 1.
Theorem Zpower_nat_1 : forall z : Z, Zpower_nat z 1 = z.
Theorem Zmin_le1 : forall z1 z2 : Z, (z1 <= z2)%Z -> Zmin z1 z2 = z1.
Theorem Zmin_le2 : forall z1 z2 : Z, (z2 <= z1)%Z -> Zmin z1 z2 = z2.
Theorem Zmin_Zle :
forall z1 z2 z3 : Z,
(z1 <= z2)%Z -> (z1 <= z3)%Z -> (z1 <= Zmin z2 z3)%Z.
Theorem Zminus_n_predm :
forall n m : Z, Zsucc (n - m) = (n - Zpred m)%Z.
Theorem Zopp_Zpred_Zs : forall z : Z, (- Zpred z)%Z = Zsucc (- z).
Theorem Zle_mult_gen :
forall x y : Z, (0 <= x)%Z -> (0 <= y)%Z -> (0 <= x * y)%Z.
Hint Resolve Zle_mult_gen: zarith.
Definition Zmax : forall x_ x_ : Z, Z :=
fun n m : Z =>
match (n ?= m)%Z with
| Datatypes.Eq => m
| Datatypes.Lt => m
| Datatypes.Gt => n
end.
Theorem ZmaxLe1 : forall z1 z2 : Z, (z1 <= Zmax z1 z2)%Z.
Theorem ZmaxSym : forall z1 z2 : Z, Zmax z1 z2 = Zmax z2 z1.
Theorem Zmax_le2 : forall z1 z2 : Z, (z1 <= z2)%Z -> Zmax z1 z2 = z2.
Theorem Zmax_le1 : forall z1 z2 : Z, (z2 <= z1)%Z -> Zmax z1 z2 = z1.
Theorem ZmaxLe2 : forall z1 z2 : Z, (z2 <= Zmax z1 z2)%Z.
Hint Resolve ZmaxLe1 ZmaxLe2: zarith.
Theorem Zeq_Zs :
forall p q : Z, (p <= q)%Z -> (q < Zsucc p)%Z -> p = q.
Theorem Zmin_Zmax : forall z1 z2 : Z, (Zmin z1 z2 <= Zmax z1 z2)%Z.
Theorem Zabs_Zmult :
forall z1 z2 : Z, Zabs (z1 * z2) = (Zabs z1 * Zabs z2)%Z.
Theorem Zle_Zmult_comp_r :
forall x y z : Z, (0 <= z)%Z -> (x <= y)%Z -> (x * z <= y * z)%Z.
Theorem Zle_Zmult_comp_l :
forall x y z : Z, (0 <= z)%Z -> (x <= y)%Z -> (z * x <= z * y)%Z.
Theorem NotZmultZero :
forall z1 z2 : Z, z1 <> 0%Z -> z2 <> 0%Z -> (z1 * z2)%Z <> 0%Z.
Hint Resolve NotZmultZero: zarith.
Theorem IZR_zero : forall p : Z, p = 0%Z -> IZR p = 0%R.
Hint Resolve not_O_INR: real.
Theorem IZR_zero_r : forall p : Z, IZR p = 0%R -> p = 0%Z.
Theorem INR_lt_nm : forall n m : nat, n < m -> (INR n < INR m)%R.
Hint Resolve INR_lt_nm: real.
Theorem Rlt_INR1 : forall n : nat, 1 < n -> (1 < INR n)%R.
Hint Resolve Rlt_INR1: real.
Theorem NEq_INR : forall n m : nat, n <> m -> INR n <> INR m.
Hint Resolve NEq_INR: real.
Theorem NEq_INRO : forall n : nat, n <> 0 -> INR n <> 0%R.
Hint Resolve NEq_INRO: real.
Theorem NEq_INR1 : forall n : nat, n <> 1 -> INR n <> 1%R.
Hint Resolve NEq_INR1: real.
Theorem not_O_lt : forall n : nat, n <> 0 -> 0 < n.
Hint Resolve not_O_lt: arith.
Theorem NEq_IZRO : forall n : Z, n <> 0%Z -> IZR n <> 0%R.
Hint Resolve NEq_IZRO: real.
Theorem Rlt_IZR : forall p q : Z, (p < q)%Z -> (IZR p < IZR q)%R.
Hint Resolve Rlt_IZR: real.
Theorem Rle_IZR : forall x y : Z, (x <= y)%Z -> (IZR x <= IZR y)%R.
Hint Resolve Rle_IZR: real.
Theorem Rlt_IZRO : forall p : Z, (0 < p)%Z -> (0 < IZR p)%R.
Hint Resolve Rlt_IZRO: real.
Theorem Rle_IZRO : forall x y : Z, (0 <= y)%Z -> (0 <= IZR y)%R.
Hint Resolve Rle_IZRO: real.
Theorem Rlt_IZR1 : forall p q : Z, (1 < q)%Z -> (1 < IZR q)%R.
Hint Resolve Rlt_IZR1: real.
Theorem Rle_IZR1 : forall x y : Z, (1 <= y)%Z -> (1 <= IZR y)%R.
Hint Resolve Rle_IZR1: real.
Theorem lt_Rlt : forall n m : nat, (INR n < INR m)%R -> n < m.
Theorem INR_inv : forall n m : nat, INR n = INR m -> n = m.
Theorem Rle_INR : forall x y : nat, x <= y -> (INR x <= INR y)%R.
Hint Resolve Rle_INR: real.
Theorem le_Rle : forall n m : nat, (INR n <= INR m)%R -> n <= m.
Theorem Rmult_IZR : forall z t : Z, IZR (z * t) = (IZR z * IZR t)%R.
Theorem absolu_Zs :
forall z : Z, (0 <= z)%Z -> Zabs_nat (Zsucc z) = S (Zabs_nat z).
Hint Resolve Zlt_le_succ: zarith.
Theorem Zlt_next :
forall n m : Z, (n < m)%Z -> m = Zsucc n \/ (Zsucc n < m)%Z.
Theorem Zle_next :
forall n m : Z, (n <= m)%Z -> m = n \/ (Zsucc n <= m)%Z.
Theorem Zlt_Zopp_Inv : forall p q : Z, (- p < - q)%Z -> (q < p)%Z.
Theorem Zle_Zopp_Inv : forall p q : Z, (- p <= - q)%Z -> (q <= p)%Z.
Theorem absolu_Zs_neg :
forall z : Z, (z < 0)%Z -> S (Zabs_nat (Zsucc z)) = Zabs_nat z.
Theorem Zlt_absolu :
forall (x : Z) (n : nat), Zabs_nat x < n -> (x < Z_of_nat n)%Z.
Theorem inj_pred :
forall n : nat, n <> 0 -> Z_of_nat (pred n) = Zpred (Z_of_nat n).
Theorem Zle_abs : forall p : Z, (p <= Z_of_nat (Zabs_nat p))%Z.
Hint Resolve Zle_abs: zarith.
Theorem ZleAbs :
forall (z : Z) (n : nat),
(- Z_of_nat n <= z)%Z -> (z <= Z_of_nat n)%Z -> Zabs_nat z <= n.
Theorem lt_Zlt_inv : forall n m : nat, (Z_of_nat n < Z_of_nat m)%Z -> n < m.
Theorem NconvertO : forall p : positive, nat_of_P p <> 0.
Hint Resolve NconvertO: zarith.
Theorem absolu_lt_nz : forall z : Z, z <> 0%Z -> 0 < Zabs_nat z.
Theorem Rlt2 : (0 < INR 2)%R.
Hint Resolve Rlt2: real.
Theorem RlIt2 : (0 < / INR 2)%R.
Hint Resolve RlIt2: real.
Theorem Rledouble : forall r : R, (0 <= r)%R -> (r <= INR 2 * r)%R.
Theorem Rltdouble : forall r : R, (0 < r)%R -> (r < INR 2 * r)%R.
Theorem Rlt_RinvDouble : forall r : R, (0 < r)%R -> (/ INR 2 * r < r)%R.
Hint Resolve Rledouble: real.
Theorem Rle_Rinv : forall x y : R, (0 < x)%R -> (x <= y)%R -> (/ y <= / x)%R.
Theorem Int_part_INR : forall n : nat, Int_part (INR n) = Z_of_nat n.
Theorem Int_part_IZR : forall z : Z, Int_part (IZR z) = z.
Theorem Zlt_Rlt : forall z1 z2 : Z, (IZR z1 < IZR z2)%R -> (z1 < z2)%Z.
Theorem Zle_Rle :
forall z1 z2 : Z, (IZR z1 <= IZR z2)%R -> (z1 <= z2)%Z.
Theorem IZR_inv : forall z1 z2 : Z, IZR z1 = IZR z2 :>R -> z1 = z2.
Theorem Zabs_eq_opp : forall x, (x <= 0)%Z -> Zabs x = (- x)%Z.
Theorem Zabs_Zs : forall z : Z, (Zabs (Zsucc z) <= Zsucc (Zabs z))%Z.
Hint Resolve Zabs_Zs: zarith.
Theorem Zle_Zpred : forall x y : Z, (x < y)%Z -> (x <= Zpred y)%Z.
Hint Resolve Zle_Zpred: zarith.
Theorem Zabs_Zopp : forall z : Z, Zabs (- z) = Zabs z.
Theorem Zle_Zabs : forall z : Z, (z <= Zabs z)%Z.
Hint Resolve Zle_Zabs: zarith.
Theorem Zlt_mult_simpl_l :
forall a b c : Z, (0 < c)%Z -> (c * a < c * b)%Z -> (a < b)%Z.
Fixpoint pos_eq_bool (a b : positive) {struct b} : bool :=
match a, b with
| xH, xH => true
| xI a', xI b' => pos_eq_bool a' b'
| xO a', xO b' => pos_eq_bool a' b'
| _, _ => false
end.
Theorem pos_eq_bool_correct :
forall p q : positive,
match pos_eq_bool p q with
| true => p = q
| false => p <> q
end.
Theorem Z_O_1 : (0 < 1)%Z.
Hint Resolve Z_O_1: zarith.
Definition Z_eq_bool a b :=
match a, b with
| Z0, Z0 => true
| Zpos a', Zpos b' => pos_eq_bool a' b'
| Zneg a', Zneg b' => pos_eq_bool a' b'
| _, _ => false
end.
Theorem Z_eq_bool_correct :
forall p q : Z,
match Z_eq_bool p q with
| true => p = q
| false => p <> q
end.
Theorem Zlt_mult_ZERO :
forall x y : Z, (0 < x)%Z -> (0 < y)%Z -> (0 < x * y)%Z.
Hint Resolve Zlt_mult_ZERO: zarith.
Theorem Zlt_Zminus_ZERO :
forall z1 z2 : Z, (z2 < z1)%Z -> (0 < z1 - z2)%Z.
Theorem Zle_Zminus_ZERO :
forall z1 z2 : Z, (z2 <= z1)%Z -> (0 <= z1 - z2)%Z.
Hint Resolve Zle_Zminus_ZERO Zlt_Zminus_ZERO: zarith.
Theorem Zle_Zpred_Zpred :
forall z1 z2 : Z, (z1 <= z2)%Z -> (Zpred z1 <= Zpred z2)%Z.
Hint Resolve Zle_Zpred_Zpred: zarith.
Theorem Zle_ZERO_Zabs : forall z : Z, (0 <= Zabs z)%Z.
Hint Resolve Zle_ZERO_Zabs: zarith.
Theorem Zlt_Zabs_inv1 :
forall z1 z2 : Z, (Zabs z1 < z2)%Z -> (- z2 < z1)%Z.
Theorem Zlt_Zabs_inv2 :
forall z1 z2 : Z, (Zabs z1 < Zabs z2)%Z -> (z1 < Zabs z2)%Z.
Theorem Zle_Zabs_inv1 :
forall z1 z2 : Z, (Zabs z1 <= z2)%Z -> (- z2 <= z1)%Z.
Theorem Zle_Zabs_inv2 :
forall z1 z2 : Z, (Zabs z1 <= z2)%Z -> (z1 <= z2)%Z.
Theorem Zlt_Zabs_Zpred :
forall z1 z2 : Z,
(Zabs z1 < z2)%Z -> z1 <> Zpred z2 -> (Zabs (Zsucc z1) < z2)%Z.
Theorem Zle_n_Zpred :
forall z1 z2 : Z, (Zpred z1 <= Zpred z2)%Z -> (z1 <= z2)%Z.
Theorem Zpred_Zopp_Zs : forall z : Z, Zpred (- z) = (- Zsucc z)%Z.
Theorem Zlt_1_O : forall z : Z, (1 <= z)%Z -> (0 < z)%Z.
Hint Resolve Zlt_succ Zsucc_lt_compat Zle_lt_succ: zarith.
Theorem Zlt_not_eq : forall p q : Z, (p < q)%Z -> p <> q.
Theorem Zlt_not_eq_rev : forall p q : Z, (q < p)%Z -> p <> q.
Hint Resolve Zlt_not_eq Zlt_not_eq_rev: zarith.
Theorem Zle_Zpred_Zlt :
forall z1 z2 : Z, (z1 <= z2)%Z -> (Zpred z1 < z2)%Z.
Hint Resolve Zle_Zpred_Zlt: zarith.
Theorem Zle_Zpred_inv :
forall z1 z2 : Z, (z1 <= Zpred z2)%Z -> (z1 < z2)%Z.
Theorem Zabs_intro :
forall (P : Z -> Prop) (z : Z), P (- z)%Z -> P z -> P (Zabs z).
Theorem Zpred_Zle_Zabs_intro :
forall z1 z2 : Z,
(- Zpred z2 <= z1)%Z -> (z1 <= Zpred z2)%Z -> (Zabs z1 < z2)%Z.
Theorem Zlt_ZERO_Zle_ONE : forall z : Z, (0 < z)%Z -> (1 <= z)%Z.
Hint Resolve Zlt_ZERO_Zle_ONE: zarith.
Theorem ptonat_def1 : forall p q, 1 < Pmult_nat p (S (S q)).
Hint Resolve ptonat_def1: arith.
Theorem lt_S_le : forall p q, p < q -> S p <= q.
Hint Resolve lt_S_le: arith.
Theorem Zlt_Zabs_intro :
forall z1 z2 : Z, (- z2 < z1)%Z -> (z1 < z2)%Z -> (Zabs z1 < z2)%Z.