diff options
Diffstat (limited to 'theories')
| -rw-r--r-- | theories/Arith/Cantor.v | 88 | ||||
| -rw-r--r-- | theories/Compat/Coq813.v | 10 | ||||
| -rw-r--r-- | theories/Logic/ChoiceFacts.v | 2 | ||||
| -rw-r--r-- | theories/Logic/ExtensionalityFacts.v | 2 | ||||
| -rw-r--r-- | theories/ZArith/BinInt.v | 2 | ||||
| -rw-r--r-- | theories/ZArith/ZArith.v | 6 | ||||
| -rw-r--r-- | theories/ZArith/Znat.v | 9 | ||||
| -rw-r--r-- | theories/ZArith/Zorder.v | 11 | ||||
| -rw-r--r-- | theories/ZArith/auxiliary.v | 8 | ||||
| -rw-r--r-- | theories/dune | 4 | ||||
| -rw-r--r-- | theories/index.mld | 3 | ||||
| -rw-r--r-- | theories/micromega/OrderedRing.v | 3 | ||||
| -rw-r--r-- | theories/micromega/ZifyInst.v | 20 | ||||
| -rw-r--r-- | theories/omega/Omega.v | 24 | ||||
| -rw-r--r-- | theories/omega/OmegaLemmas.v | 44 | ||||
| -rw-r--r-- | theories/omega/OmegaPlugin.v | 17 | ||||
| -rw-r--r-- | theories/omega/OmegaTactic.v | 17 | ||||
| -rw-r--r-- | theories/omega/PreOmega.v | 448 |
18 files changed, 124 insertions, 594 deletions
diff --git a/theories/Arith/Cantor.v b/theories/Arith/Cantor.v new file mode 100644 index 0000000000..b63d970db7 --- /dev/null +++ b/theories/Arith/Cantor.v @@ -0,0 +1,88 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * Copyright INRIA, CNRS and contributors *) +(* <O___,, * (see version control and CREDITS file for authors & dates) *) +(* \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) *) +(************************************************************************) + +(** Implementation of the Cantor pairing and its inverse function *) + +Require Import PeanoNat Lia. + +(** Bijections between [nat * nat] and [nat] *) + +(** Cantor pairing [to_nat] *) + +Definition to_nat '(x, y) : nat := + y + (nat_rec _ 0 (fun i m => (S i) + m) (y + x)). + +(** Cantor pairing inverse [of_nat] *) + +Definition of_nat (n : nat) : nat * nat := + nat_rec _ (0, 0) (fun _ '(x, y) => + match x with | S x => (x, S y) | _ => (S y, 0) end) n. + +(** [of_nat] is the left inverse for [to_nat] *) + +Lemma cancel_of_to p : of_nat (to_nat p) = p. +Proof. + enough (H : forall n p, to_nat p = n -> of_nat n = p) by now apply H. + intro n. induction n as [|n IHn]. + - now intros [[|?] [|?]]. + - intros [x [|y]]. + + destruct x as [|x]; [discriminate|]. + intros [=H]. cbn. fold (of_nat n). + rewrite (IHn (0, x)); [reflexivity|]. + rewrite <- H. cbn. now rewrite PeanoNat.Nat.add_0_r. + + intros [=H]. cbn. fold (of_nat n). + rewrite (IHn (S x, y)); [reflexivity|]. + rewrite <- H. cbn. now rewrite Nat.add_succ_r. +Qed. + +(** [to_nat] is injective *) + +Corollary to_nat_inj p q : to_nat p = to_nat q -> p = q. +Proof. + intros H %(f_equal of_nat). now rewrite ?cancel_of_to in H. +Qed. + +(** [to_nat] is the left inverse for [of_nat] *) + +Lemma cancel_to_of n : to_nat (of_nat n) = n. +Proof. + induction n as [|n IHn]; [reflexivity|]. + cbn. fold (of_nat n). destruct (of_nat n) as [[|x] y]. + - rewrite <- IHn. cbn. now rewrite PeanoNat.Nat.add_0_r. + - rewrite <- IHn. cbn. now rewrite (Nat.add_succ_r y x). +Qed. + +(** [of_nat] is injective *) + +Corollary of_nat_inj n m : of_nat n = of_nat m -> n = m. +Proof. + intros H %(f_equal to_nat). now rewrite ?cancel_to_of in H. +Qed. + +(** Polynomial specifications of [to_nat] *) + +Lemma to_nat_spec x y : + to_nat (x, y) * 2 = y * 2 + (y + x) * S (y + x). +Proof. + cbn. induction (y + x) as [|n IHn]; cbn; lia. +Qed. + +Lemma to_nat_spec2 x y : + to_nat (x, y) = y + (y + x) * S (y + x) / 2. +Proof. + now rewrite <- Nat.div_add_l, <- to_nat_spec, Nat.div_mul. +Qed. + +(** [to_nat] is non-decreasing in (the sum of) pair components *) + +Lemma to_nat_non_decreasing x y : y + x <= to_nat (x, y). +Proof. + pose proof (to_nat_spec x y). nia. +Qed. diff --git a/theories/Compat/Coq813.v b/theories/Compat/Coq813.v index fe7431dcd3..5cfb9d84c7 100644 --- a/theories/Compat/Coq813.v +++ b/theories/Compat/Coq813.v @@ -11,3 +11,13 @@ (** Compatibility file for making Coq act similar to Coq v8.13 *) Require Export Coq.Compat.Coq814. + +Require Coq.micromega.Lia. +Module Export Coq. + Module Export omega. + Module Export Omega. + #[deprecated(since="8.12", note="The omega tactic was removed in v8.14. You're now relying on the lia tactic.")] + Ltac omega := Lia.lia. + End Omega. + End omega. +End Coq. diff --git a/theories/Logic/ChoiceFacts.v b/theories/Logic/ChoiceFacts.v index 3dac62c476..23bc396fd7 100644 --- a/theories/Logic/ChoiceFacts.v +++ b/theories/Logic/ChoiceFacts.v @@ -676,7 +676,7 @@ Qed. We show instead that functional relation reification and the functional form of the axiom of choice are equivalent on decidable - relation with [nat] as codomain + relations with [nat] as codomain. *) Require Import Wf_nat. diff --git a/theories/Logic/ExtensionalityFacts.v b/theories/Logic/ExtensionalityFacts.v index 4735d6c9ca..a151cca3af 100644 --- a/theories/Logic/ExtensionalityFacts.v +++ b/theories/Logic/ExtensionalityFacts.v @@ -43,7 +43,7 @@ Definition is_inverse A B f g := (forall a:A, g (f a) = a) /\ (forall b:B, f (g #[universes(template)] Record Delta A := { pi1:A; pi2:A; eq:pi1=pi2 }. -Definition delta {A} (a:A) := {|pi1 := a; pi2 := a; eq := eq_refl a |}. +Definition delta {A} (a:A) := {| pi1 := a; pi2 := a; eq := eq_refl a |}. Arguments pi1 {A} _. Arguments pi2 {A} _. diff --git a/theories/ZArith/BinInt.v b/theories/ZArith/BinInt.v index 47137414dc..ff1a2b5a53 100644 --- a/theories/ZArith/BinInt.v +++ b/theories/ZArith/BinInt.v @@ -1775,8 +1775,6 @@ weak_Zmult_plus_distr_r (now Z.mul_add_distr_pos) Definition Zne (x y:Z) := x <> y. (* TODO : to remove someday ? *) -Register Zne as plugins.omega.Zne. - Ltac elim_compare com1 com2 := case (Dcompare (com1 ?= com2)%Z); [ idtac | let x := fresh "H" in diff --git a/theories/ZArith/ZArith.v b/theories/ZArith/ZArith.v index ed47539e1e..5490044ac7 100644 --- a/theories/ZArith/ZArith.v +++ b/theories/ZArith/ZArith.v @@ -16,9 +16,11 @@ Require Export ZArith_base. Require Export Zpow_def. -(** Extra modules using [Omega] or [Ring]. *) +(** Extra modules using [Ring]. *) -Require Export Omega. +Require Export OmegaLemmas. +Require Export PreOmega. +Require Export ZArith_hints. Require Export Zcomplements. Require Export Zpower. Require Export Zdiv. diff --git a/theories/ZArith/Znat.v b/theories/ZArith/Znat.v index 7f72d42d1f..b7cd849323 100644 --- a/theories/ZArith/Znat.v +++ b/theories/ZArith/Znat.v @@ -959,13 +959,6 @@ Definition inj_lt n m := proj1 (Nat2Z.inj_lt n m). Definition inj_ge n m := proj1 (Nat2Z.inj_ge n m). Definition inj_gt n m := proj1 (Nat2Z.inj_gt n m). -Register neq as plugins.omega.neq. -Register inj_eq as plugins.omega.inj_eq. -Register inj_neq as plugins.omega.inj_neq. -Register inj_le as plugins.omega.inj_le. -Register inj_lt as plugins.omega.inj_lt. -Register inj_ge as plugins.omega.inj_ge. -Register inj_gt as plugins.omega.inj_gt. (** For the others, a Notation is fine *) @@ -1033,5 +1026,3 @@ Theorem inj_minus2 : forall n m:nat, (m > n)%nat -> Z.of_nat (n - m) = 0. Proof. intros. rewrite not_le_minus_0; auto with arith. Qed. - -Register inj_minus2 as plugins.omega.inj_minus2. diff --git a/theories/ZArith/Zorder.v b/theories/ZArith/Zorder.v index 4c533ac458..bef9cede12 100644 --- a/theories/ZArith/Zorder.v +++ b/theories/ZArith/Zorder.v @@ -64,11 +64,6 @@ Proof. apply Z.lt_gt_cases. Qed. -Register dec_Zne as plugins.omega.dec_Zne. -Register dec_Zgt as plugins.omega.dec_Zgt. -Register dec_Zge as plugins.omega.dec_Zge. -Register not_Zeq as plugins.omega.not_Zeq. - (** * Relating strict and large orders *) Notation Zgt_iff_lt := Z.gt_lt_iff (only parsing). @@ -120,12 +115,6 @@ Proof. destruct (Z.eq_decidable n m); [assumption|now elim H]. Qed. -Register Znot_le_gt as plugins.omega.Znot_le_gt. -Register Znot_lt_ge as plugins.omega.Znot_lt_ge. -Register Znot_ge_lt as plugins.omega.Znot_ge_lt. -Register Znot_gt_le as plugins.omega.Znot_gt_le. -Register not_Zne as plugins.omega.not_Zne. - (** * Equivalence and order properties *) (** Reflexivity *) diff --git a/theories/ZArith/auxiliary.v b/theories/ZArith/auxiliary.v index 10ea6cc03e..369a0c46ae 100644 --- a/theories/ZArith/auxiliary.v +++ b/theories/ZArith/auxiliary.v @@ -93,11 +93,3 @@ Proof. apply Z.le_lt_trans with (m*n+p); trivial. now apply Z.add_lt_mono_l. Qed. - -Register Zegal_left as plugins.omega.Zegal_left. -Register Zne_left as plugins.omega.Zne_left. -Register Zlt_left as plugins.omega.Zlt_left. -Register Zgt_left as plugins.omega.Zgt_left. -Register Zle_left as plugins.omega.Zle_left. -Register Zge_left as plugins.omega.Zge_left. -Register Zmult_le_approx as plugins.omega.Zmult_le_approx. diff --git a/theories/dune b/theories/dune index 1cd3d8c119..57b97f080c 100644 --- a/theories/dune +++ b/theories/dune @@ -22,7 +22,6 @@ coq-core.plugins.ring coq-core.plugins.nsatz - coq-core.plugins.omega coq-core.plugins.zify coq-core.plugins.micromega @@ -34,3 +33,6 @@ coq-core.plugins.derive)) (include_subdirs qualified) + +(documentation + (package coq-stdlib)) diff --git a/theories/index.mld b/theories/index.mld new file mode 100644 index 0000000000..360864342b --- /dev/null +++ b/theories/index.mld @@ -0,0 +1,3 @@ +{0 coq-stdlib } + +The coq-stdlib package only contains Coq theory files for the standard library and no OCaml libraries. diff --git a/theories/micromega/OrderedRing.v b/theories/micromega/OrderedRing.v index 5fa3740ab1..bfbd6fd8d3 100644 --- a/theories/micromega/OrderedRing.v +++ b/theories/micromega/OrderedRing.v @@ -423,7 +423,7 @@ Qed. (* The following theorems are used to build a morphism from Z to R and prove its properties in ZCoeff.v. They are not used in RingMicromega.v. *) -(* Surprisingly, multilication is needed to prove the following theorem *) +(* Surprisingly, multiplication is needed to prove the following theorem *) Theorem Ropp_neg_pos : forall n : R, - n < 0 <-> 0 < n. Proof. @@ -457,4 +457,3 @@ apply Rtimes_pos_pos. assumption. now apply -> Rlt_lt_minus. Qed.*) End STRICT_ORDERED_RING. - diff --git a/theories/micromega/ZifyInst.v b/theories/micromega/ZifyInst.v index 9881e73f76..dd31b998d4 100644 --- a/theories/micromega/ZifyInst.v +++ b/theories/micromega/ZifyInst.v @@ -307,15 +307,15 @@ Instance Op_N_mul : BinOp N.mul := Add Zify BinOp Op_N_mul. Instance Op_N_sub : BinOp N.sub := - {| TBOp := fun x y => Z.max 0 (x - y) ; TBOpInj := N2Z.inj_sub_max|}. + {| TBOp := fun x y => Z.max 0 (x - y) ; TBOpInj := N2Z.inj_sub_max |}. Add Zify BinOp Op_N_sub. Instance Op_N_div : BinOp N.div := - {| TBOp := Z.div ; TBOpInj := N2Z.inj_div|}. + {| TBOp := Z.div ; TBOpInj := N2Z.inj_div |}. Add Zify BinOp Op_N_div. Instance Op_N_mod : BinOp N.modulo := - {| TBOp := Z.rem ; TBOpInj := N2Z.inj_rem|}. + {| TBOp := Z.rem ; TBOpInj := N2Z.inj_rem |}. Add Zify BinOp Op_N_mod. Instance Op_N_pred : UnOp N.pred := @@ -332,19 +332,19 @@ Add Zify UnOp Op_N_succ. (* zify_Z_rel *) Instance Op_Z_ge : BinRel Z.ge := - {| TR := Z.ge ; TRInj := fun x y => iff_refl (x>= y)|}. + {| TR := Z.ge ; TRInj := fun x y => iff_refl (x>= y) |}. Add Zify BinRel Op_Z_ge. Instance Op_Z_lt : BinRel Z.lt := - {| TR := Z.lt ; TRInj := fun x y => iff_refl (x < y)|}. + {| TR := Z.lt ; TRInj := fun x y => iff_refl (x < y) |}. Add Zify BinRel Op_Z_lt. Instance Op_Z_gt : BinRel Z.gt := - {| TR := Z.gt ;TRInj := fun x y => iff_refl (x > y)|}. + {| TR := Z.gt ;TRInj := fun x y => iff_refl (x > y) |}. Add Zify BinRel Op_Z_gt. Instance Op_Z_le : BinRel Z.le := - {| TR := Z.le ;TRInj := fun x y => iff_refl (x <= y)|}. + {| TR := Z.le ;TRInj := fun x y => iff_refl (x <= y) |}. Add Zify BinRel Op_Z_le. Instance Op_eqZ : BinRel (@eq Z) := @@ -460,7 +460,7 @@ Add Zify UnOp Op_Z_to_nat. (** Specification of derived operators over Z *) Instance ZmaxSpec : BinOpSpec Z.max := - {| BPred := fun n m r => n < m /\ r = m \/ m <= n /\ r = n ; BSpec := Z.max_spec|}. + {| BPred := fun n m r => n < m /\ r = m \/ m <= n /\ r = n ; BSpec := Z.max_spec |}. Add Zify BinOpSpec ZmaxSpec. Instance ZminSpec : BinOpSpec Z.min := @@ -470,12 +470,12 @@ Add Zify BinOpSpec ZminSpec. Instance ZsgnSpec : UnOpSpec Z.sgn := {| UPred := fun n r : Z => 0 < n /\ r = 1 \/ 0 = n /\ r = 0 \/ n < 0 /\ r = - (1) ; - USpec := Z.sgn_spec|}. + USpec := Z.sgn_spec |}. Add Zify UnOpSpec ZsgnSpec. Instance ZabsSpec : UnOpSpec Z.abs := {| UPred := fun n r: Z => 0 <= n /\ r = n \/ n < 0 /\ r = - n ; - USpec := Z.abs_spec|}. + USpec := Z.abs_spec |}. Add Zify UnOpSpec ZabsSpec. (** Saturate positivity constraints *) diff --git a/theories/omega/Omega.v b/theories/omega/Omega.v deleted file mode 100644 index 5c52284621..0000000000 --- a/theories/omega/Omega.v +++ /dev/null @@ -1,24 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * Copyright INRIA, CNRS and contributors *) -(* <O___,, * (see version control and CREDITS file for authors & dates) *) -(* \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) *) -(************************************************************************) -(**************************************************************************) -(* *) -(* Omega: a solver of quantifier-free problems in Presburger Arithmetic *) -(* *) -(* Pierre Crégut (CNET, Lannion, France) *) -(* *) -(**************************************************************************) - -(* We import what is necessary for Omega *) -Require Export ZArith_base. -Require Export OmegaLemmas. -Require Export PreOmega. -Require Export ZArith_hints. - -Declare ML Module "omega_plugin". diff --git a/theories/omega/OmegaLemmas.v b/theories/omega/OmegaLemmas.v index 08e9ac345d..8ccd8d76f6 100644 --- a/theories/omega/OmegaLemmas.v +++ b/theories/omega/OmegaLemmas.v @@ -261,47 +261,3 @@ Proof. intros n; exists (Z.of_nat n); split; trivial. rewrite Z.mul_1_r, Z.add_0_r. apply Nat2Z.is_nonneg. Qed. - -Register fast_Zplus_assoc_reverse as plugins.omega.fast_Zplus_assoc_reverse. -Register fast_Zplus_assoc as plugins.omega.fast_Zplus_assoc. -Register fast_Zmult_assoc_reverse as plugins.omega.fast_Zmult_assoc_reverse. -Register fast_Zplus_permute as plugins.omega.fast_Zplus_permute. -Register fast_Zplus_comm as plugins.omega.fast_Zplus_comm. -Register fast_Zmult_comm as plugins.omega.fast_Zmult_comm. - -Register OMEGA1 as plugins.omega.OMEGA1. -Register OMEGA2 as plugins.omega.OMEGA2. -Register OMEGA3 as plugins.omega.OMEGA3. -Register OMEGA4 as plugins.omega.OMEGA4. -Register OMEGA5 as plugins.omega.OMEGA5. -Register OMEGA6 as plugins.omega.OMEGA6. -Register OMEGA7 as plugins.omega.OMEGA7. -Register OMEGA8 as plugins.omega.OMEGA8. -Register OMEGA9 as plugins.omega.OMEGA9. -Register fast_OMEGA10 as plugins.omega.fast_OMEGA10. -Register fast_OMEGA11 as plugins.omega.fast_OMEGA11. -Register fast_OMEGA12 as plugins.omega.fast_OMEGA12. -Register fast_OMEGA13 as plugins.omega.fast_OMEGA13. -Register fast_OMEGA14 as plugins.omega.fast_OMEGA14. -Register fast_OMEGA15 as plugins.omega.fast_OMEGA15. -Register fast_OMEGA16 as plugins.omega.fast_OMEGA16. -Register OMEGA17 as plugins.omega.OMEGA17. -Register OMEGA18 as plugins.omega.OMEGA18. -Register OMEGA19 as plugins.omega.OMEGA19. -Register OMEGA20 as plugins.omega.OMEGA20. - -Register fast_Zred_factor0 as plugins.omega.fast_Zred_factor0. -Register fast_Zred_factor1 as plugins.omega.fast_Zred_factor1. -Register fast_Zred_factor2 as plugins.omega.fast_Zred_factor2. -Register fast_Zred_factor3 as plugins.omega.fast_Zred_factor3. -Register fast_Zred_factor4 as plugins.omega.fast_Zred_factor4. -Register fast_Zred_factor5 as plugins.omega.fast_Zred_factor5. -Register fast_Zred_factor6 as plugins.omega.fast_Zred_factor6. - -Register fast_Zmult_plus_distr_l as plugins.omega.fast_Zmult_plus_distr_l. -Register fast_Zopp_plus_distr as plugins.omega.fast_Zopp_plus_distr. -Register fast_Zopp_mult_distr_r as plugins.omega.fast_Zopp_mult_distr_r. -Register fast_Zopp_eq_mult_neg_1 as plugins.omega.fast_Zopp_eq_mult_neg_1. - -Register new_var as plugins.omega.new_var. -Register intro_Z as plugins.omega.intro_Z. diff --git a/theories/omega/OmegaPlugin.v b/theories/omega/OmegaPlugin.v deleted file mode 100644 index e0cf24f6aa..0000000000 --- a/theories/omega/OmegaPlugin.v +++ /dev/null @@ -1,17 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * Copyright INRIA, CNRS and contributors *) -(* <O___,, * (see version control and CREDITS file for authors & dates) *) -(* \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) *) -(************************************************************************) - -(* To strictly import the omega tactic *) - -Require ZArith_base. -Require OmegaLemmas. -Require PreOmega. - -Declare ML Module "omega_plugin". diff --git a/theories/omega/OmegaTactic.v b/theories/omega/OmegaTactic.v deleted file mode 100644 index e0cf24f6aa..0000000000 --- a/theories/omega/OmegaTactic.v +++ /dev/null @@ -1,17 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * Copyright INRIA, CNRS and contributors *) -(* <O___,, * (see version control and CREDITS file for authors & dates) *) -(* \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) *) -(************************************************************************) - -(* To strictly import the omega tactic *) - -Require ZArith_base. -Require OmegaLemmas. -Require PreOmega. - -Declare ML Module "omega_plugin". diff --git a/theories/omega/PreOmega.v b/theories/omega/PreOmega.v index 70f25e7243..bf873785d0 100644 --- a/theories/omega/PreOmega.v +++ b/theories/omega/PreOmega.v @@ -12,9 +12,10 @@ Require Import Arith Max Min BinInt BinNat Znat Nnat. Local Open Scope Z_scope. -(** * [Z.div_mod_to_equations], [Z.quot_rem_to_equations], [Z.to_euclidean_division_equations]: the tactics for preprocessing [Z.div] and [Z.modulo], [Z.quot] and [Z.rem] *) +(** * [Z.div_mod_to_equations], [Z.quot_rem_to_equations], [Z.to_euclidean_division_equations]: + the tactics for preprocessing [Z.div] and [Z.modulo], [Z.quot] and [Z.rem] *) -(** These tactic use the complete specification of [Z.div] and +(** These tactics use the complete specification of [Z.div] and [Z.modulo] ([Z.quot] and [Z.rem], respectively) to remove these functions from the goal without losing information. The [Z.euclidean_division_equations_cleanup] tactic removes needless @@ -127,449 +128,6 @@ Module Z. Ltac to_euclidean_division_equations := div_mod_to_equations'; quot_rem_to_equations'; euclidean_division_equations_cleanup. End Z. -Set Warnings "-deprecated-tactic". - -(** * zify: the Z-ification tactic *) - -(* This tactic searches for nat and N and positive elements in the goal and - translates everything into Z. It is meant as a pre-processor for - (r)omega; for instance a positivity hypothesis is added whenever - - a multiplication is encountered - - an atom is encountered (that is a variable or an unknown construct) - - Recognized relations (can be handled as deeply as allowed by setoid rewrite): - - { eq, le, lt, ge, gt } on { Z, positive, N, nat } - - Recognized operations: - - on Z: Z.min, Z.max, Z.abs, Z.sgn are translated in term of <= < = - - on nat: + * - S O pred min max Pos.to_nat N.to_nat Z.abs_nat - - on positive: Zneg Zpos xI xO xH + * - Pos.succ Pos.pred Pos.min Pos.max Pos.of_succ_nat - - on N: N0 Npos + * - N.pred N.succ N.min N.max N.of_nat Z.abs_N -*) - - - - -(** I) translation of Z.max, Z.min, Z.abs, Z.sgn into recognized equations *) - -#[deprecated( note = "Use 'zify' instead")] -Ltac zify_unop_core t thm a := - (* Let's introduce the specification theorem for t *) - pose proof (thm a); - (* Then we replace (t a) everywhere with a fresh variable *) - let z := fresh "z" in set (z:=t a) in *; clearbody z. - -#[deprecated( note = "Use 'zify' instead")] -Ltac zify_unop_var_or_term t thm a := - (* If a is a variable, no need for aliasing *) - let za := fresh "z" in - (rename a into za; rename za into a; zify_unop_core t thm a) || - (* Otherwise, a is a complex term: we alias it. *) - (remember a as za; zify_unop_core t thm za). - -#[deprecated( note = "Use 'zify' instead")] -Ltac zify_unop t thm a := - (* If a is a scalar, we can simply reduce the unop. *) - (* Note that simpl wasn't enough to reduce [Z.max 0 0] (#5439) *) - let isz := isZcst a in - match isz with - | true => - let u := eval compute in (t a) in - change (t a) with u in * - | _ => zify_unop_var_or_term t thm a - end. - -#[deprecated( note = "Use 'zify' instead")] -Ltac zify_unop_nored t thm a := - (* in this version, we don't try to reduce the unop (that can be (Z.add x)) *) - let isz := isZcst a in - match isz with - | true => zify_unop_core t thm a - | _ => zify_unop_var_or_term t thm a - end. - -#[deprecated( note = "Use 'zify' instead")] -Ltac zify_binop t thm a b:= - (* works as zify_unop, except that we should be careful when - dealing with b, since it can be equal to a *) - let isza := isZcst a in - match isza with - | true => zify_unop (t a) (thm a) b - | _ => - let za := fresh "z" in - (rename a into za; rename za into a; zify_unop_nored (t a) (thm a) b) || - (remember a as za; match goal with - | H : za = b |- _ => zify_unop_nored (t za) (thm za) za - | _ => zify_unop_nored (t za) (thm za) b - end) - end. - -#[deprecated( note = "Use 'zify' instead")] -Ltac zify_op_1 := - match goal with - | x := ?t : Z |- _ => let h := fresh "heq_" x in pose proof (eq_refl : x = t) as h; clearbody x - | |- context [ Z.max ?a ?b ] => zify_binop Z.max Z.max_spec a b - | H : context [ Z.max ?a ?b ] |- _ => zify_binop Z.max Z.max_spec a b - | |- context [ Z.min ?a ?b ] => zify_binop Z.min Z.min_spec a b - | H : context [ Z.min ?a ?b ] |- _ => zify_binop Z.min Z.min_spec a b - | |- context [ Z.sgn ?a ] => zify_unop Z.sgn Z.sgn_spec a - | H : context [ Z.sgn ?a ] |- _ => zify_unop Z.sgn Z.sgn_spec a - | |- context [ Z.abs ?a ] => zify_unop Z.abs Z.abs_spec a - | H : context [ Z.abs ?a ] |- _ => zify_unop Z.abs Z.abs_spec a - end. - -Ltac zify_op := repeat zify_op_1. - - -(** II) Conversion from nat to Z *) - - -Definition Z_of_nat' := Z.of_nat. - -Ltac hide_Z_of_nat t := - let z := fresh "z" in set (z:=Z.of_nat t) in *; - change Z.of_nat with Z_of_nat' in z; - unfold z in *; clear z. - -#[deprecated( note = "Use 'zify' instead")] -Ltac zify_nat_rel := - match goal with - (* I: equalities *) - | x := ?t : nat |- _ => let h := fresh "heq_" x in pose proof (eq_refl : x = t) as h; clearbody x - | |- (@eq nat ?a ?b) => apply (Nat2Z.inj a b) (* shortcut *) - | H : context [ @eq nat ?a ?b ] |- _ => rewrite <- (Nat2Z.inj_iff a b) in H - | |- context [ @eq nat ?a ?b ] => rewrite <- (Nat2Z.inj_iff a b) - (* II: less than *) - | H : context [ lt ?a ?b ] |- _ => rewrite (Nat2Z.inj_lt a b) in H - | |- context [ lt ?a ?b ] => rewrite (Nat2Z.inj_lt a b) - (* III: less or equal *) - | H : context [ le ?a ?b ] |- _ => rewrite (Nat2Z.inj_le a b) in H - | |- context [ le ?a ?b ] => rewrite (Nat2Z.inj_le a b) - (* IV: greater than *) - | H : context [ gt ?a ?b ] |- _ => rewrite (Nat2Z.inj_gt a b) in H - | |- context [ gt ?a ?b ] => rewrite (Nat2Z.inj_gt a b) - (* V: greater or equal *) - | H : context [ ge ?a ?b ] |- _ => rewrite (Nat2Z.inj_ge a b) in H - | |- context [ ge ?a ?b ] => rewrite (Nat2Z.inj_ge a b) - end. - -Ltac zify_nat_op := - match goal with - (* misc type conversions: positive/N/Z to nat *) - | H : context [ Z.of_nat (Pos.to_nat ?a) ] |- _ => rewrite (positive_nat_Z a) in H - | |- context [ Z.of_nat (Pos.to_nat ?a) ] => rewrite (positive_nat_Z a) - | H : context [ Z.of_nat (N.to_nat ?a) ] |- _ => rewrite (N_nat_Z a) in H - | |- context [ Z.of_nat (N.to_nat ?a) ] => rewrite (N_nat_Z a) - | H : context [ Z.of_nat (Z.abs_nat ?a) ] |- _ => rewrite (Zabs2Nat.id_abs a) in H - | |- context [ Z.of_nat (Z.abs_nat ?a) ] => rewrite (Zabs2Nat.id_abs a) - - (* plus -> Z.add *) - | H : context [ Z.of_nat (plus ?a ?b) ] |- _ => rewrite (Nat2Z.inj_add a b) in H - | |- context [ Z.of_nat (plus ?a ?b) ] => rewrite (Nat2Z.inj_add a b) - - (* min -> Z.min *) - | H : context [ Z.of_nat (min ?a ?b) ] |- _ => rewrite (Nat2Z.inj_min a b) in H - | |- context [ Z.of_nat (min ?a ?b) ] => rewrite (Nat2Z.inj_min a b) - - (* max -> Z.max *) - | H : context [ Z.of_nat (max ?a ?b) ] |- _ => rewrite (Nat2Z.inj_max a b) in H - | |- context [ Z.of_nat (max ?a ?b) ] => rewrite (Nat2Z.inj_max a b) - - (* minus -> Z.max (Z.sub ... ...) 0 *) - | H : context [ Z.of_nat (minus ?a ?b) ] |- _ => rewrite (Nat2Z.inj_sub_max a b) in H - | |- context [ Z.of_nat (minus ?a ?b) ] => rewrite (Nat2Z.inj_sub_max a b) - - (* pred -> minus ... -1 -> Z.max (Z.sub ... -1) 0 *) - | H : context [ Z.of_nat (pred ?a) ] |- _ => rewrite (pred_of_minus a) in H - | |- context [ Z.of_nat (pred ?a) ] => rewrite (pred_of_minus a) - - (* mult -> Z.mul and a positivity hypothesis *) - | H : context [ Z.of_nat (mult ?a ?b) ] |- _ => - pose proof (Nat2Z.is_nonneg (mult a b)); - rewrite (Nat2Z.inj_mul a b) in * - | |- context [ Z.of_nat (mult ?a ?b) ] => - pose proof (Nat2Z.is_nonneg (mult a b)); - rewrite (Nat2Z.inj_mul a b) in * - - (* O -> Z0 *) - | H : context [ Z.of_nat O ] |- _ => change (Z.of_nat O) with Z0 in H - | |- context [ Z.of_nat O ] => change (Z.of_nat O) with Z0 - - (* S -> number or Z.succ *) - | H : context [ Z.of_nat (S ?a) ] |- _ => - let isnat := isnatcst a in - match isnat with - | true => - let t := eval compute in (Z.of_nat (S a)) in - change (Z.of_nat (S a)) with t in H - | _ => rewrite (Nat2Z.inj_succ a) in H - | _ => (* if the [rewrite] fails (most likely a dependent occurrence of [Z.of_nat (S a)]), - hide [Z.of_nat (S a)] in this one hypothesis *) - change (Z.of_nat (S a)) with (Z_of_nat' (S a)) in H - end - | |- context [ Z.of_nat (S ?a) ] => - let isnat := isnatcst a in - match isnat with - | true => - let t := eval compute in (Z.of_nat (S a)) in - change (Z.of_nat (S a)) with t - | _ => rewrite (Nat2Z.inj_succ a) - | _ => (* if the [rewrite] fails (most likely a dependent occurrence of [Z.of_nat (S a)]), - hide [Z.of_nat (S a)] in the goal *) - change (Z.of_nat (S a)) with (Z_of_nat' (S a)) - end - - (* atoms of type nat : we add a positivity condition (if not already there) *) - | _ : 0 <= Z.of_nat ?a |- _ => hide_Z_of_nat a - | _ : context [ Z.of_nat ?a ] |- _ => - pose proof (Nat2Z.is_nonneg a); hide_Z_of_nat a - | |- context [ Z.of_nat ?a ] => - pose proof (Nat2Z.is_nonneg a); hide_Z_of_nat a - end. - -#[deprecated( note = "Use 'zify' instead")] -Ltac zify_nat := repeat zify_nat_rel; repeat zify_nat_op; unfold Z_of_nat' in *. - -(* III) conversion from positive to Z *) - -Definition Zpos' := Zpos. -Definition Zneg' := Zneg. - -Ltac hide_Zpos t := - let z := fresh "z" in set (z:=Zpos t) in *; - change Zpos with Zpos' in z; - unfold z in *; clear z. - -#[deprecated( note = "Use 'zify' instead")] -Ltac zify_positive_rel := - match goal with - (* I: equalities *) - | x := ?t : positive |- _ => let h := fresh "heq_" x in pose proof (eq_refl : x = t) as h; clearbody x - | |- (@eq positive ?a ?b) => apply Pos2Z.inj - | H : context [ @eq positive ?a ?b ] |- _ => rewrite <- (Pos2Z.inj_iff a b) in H - | |- context [ @eq positive ?a ?b ] => rewrite <- (Pos2Z.inj_iff a b) - (* II: less than *) - | H : context [ (?a < ?b)%positive ] |- _ => change (a<b)%positive with (Zpos a<Zpos b) in H - | |- context [ (?a < ?b)%positive ] => change (a<b)%positive with (Zpos a<Zpos b) - (* III: less or equal *) - | H : context [ (?a <= ?b)%positive ] |- _ => change (a<=b)%positive with (Zpos a<=Zpos b) in H - | |- context [ (?a <= ?b)%positive ] => change (a<=b)%positive with (Zpos a<=Zpos b) - (* IV: greater than *) - | H : context [ (?a > ?b)%positive ] |- _ => change (a>b)%positive with (Zpos a>Zpos b) in H - | |- context [ (?a > ?b)%positive ] => change (a>b)%positive with (Zpos a>Zpos b) - (* V: greater or equal *) - | H : context [ (?a >= ?b)%positive ] |- _ => change (a>=b)%positive with (Zpos a>=Zpos b) in H - | |- context [ (?a >= ?b)%positive ] => change (a>=b)%positive with (Zpos a>=Zpos b) - end. - -#[deprecated( note = "Use 'zify' instead")] -Ltac zify_positive_op := - match goal with - (* Z.pow_pos -> Z.pow *) - | H : context [ Z.pow_pos ?a ?b ] |- _ => change (Z.pow_pos a b) with (Z.pow a (Z.pos b)) in H - | |- context [ Z.pow_pos ?a ?b ] => change (Z.pow_pos a b) with (Z.pow a (Z.pos b)) - (* Zneg -> -Zpos (except for numbers) *) - | H : context [ Zneg ?a ] |- _ => - let isp := isPcst a in - match isp with - | true => change (Zneg a) with (Zneg' a) in H - | _ => change (Zneg a) with (- Zpos a) in H - end - | |- context [ Zneg ?a ] => - let isp := isPcst a in - match isp with - | true => change (Zneg a) with (Zneg' a) - | _ => change (Zneg a) with (- Zpos a) - end - - (* misc type conversions: nat to positive *) - | H : context [ Zpos (Pos.of_succ_nat ?a) ] |- _ => rewrite (Zpos_P_of_succ_nat a) in H - | |- context [ Zpos (Pos.of_succ_nat ?a) ] => rewrite (Zpos_P_of_succ_nat a) - - (* Z.power_pos *) - | H : context [ Zpos (Pos.of_succ_nat ?a) ] |- _ => rewrite (Zpos_P_of_succ_nat a) in H - | |- context [ Zpos (Pos.of_succ_nat ?a) ] => rewrite (Zpos_P_of_succ_nat a) - - (* Pos.add -> Z.add *) - | H : context [ Zpos (?a + ?b) ] |- _ => change (Zpos (a+b)) with (Zpos a + Zpos b) in H - | |- context [ Zpos (?a + ?b) ] => change (Zpos (a+b)) with (Zpos a + Zpos b) - - (* Pos.min -> Z.min *) - | H : context [ Zpos (Pos.min ?a ?b) ] |- _ => rewrite (Pos2Z.inj_min a b) in H - | |- context [ Zpos (Pos.min ?a ?b) ] => rewrite (Pos2Z.inj_min a b) - - (* Pos.max -> Z.max *) - | H : context [ Zpos (Pos.max ?a ?b) ] |- _ => rewrite (Pos2Z.inj_max a b) in H - | |- context [ Zpos (Pos.max ?a ?b) ] => rewrite (Pos2Z.inj_max a b) - - (* Pos.sub -> Z.max 1 (Z.sub ... ...) *) - | H : context [ Zpos (Pos.sub ?a ?b) ] |- _ => rewrite (Pos2Z.inj_sub_max a b) in H - | |- context [ Zpos (Pos.sub ?a ?b) ] => rewrite (Pos2Z.inj_sub_max a b) - - (* Pos.succ -> Z.succ *) - | H : context [ Zpos (Pos.succ ?a) ] |- _ => rewrite (Pos2Z.inj_succ a) in H - | |- context [ Zpos (Pos.succ ?a) ] => rewrite (Pos2Z.inj_succ a) - - (* Pos.pred -> Pos.sub ... -1 -> Z.max 1 (Z.sub ... - 1) *) - | H : context [ Zpos (Pos.pred ?a) ] |- _ => rewrite <- (Pos.sub_1_r a) in H - | |- context [ Zpos (Pos.pred ?a) ] => rewrite <- (Pos.sub_1_r a) - - (* Pos.mul -> Z.mul and a positivity hypothesis *) - | H : context [ Zpos (?a * ?b) ] |- _ => - pose proof (Pos2Z.is_pos (Pos.mul a b)); - change (Zpos (a*b)) with (Zpos a * Zpos b) in * - | |- context [ Zpos (?a * ?b) ] => - pose proof (Pos2Z.is_pos (Pos.mul a b)); - change (Zpos (a*b)) with (Zpos a * Zpos b) in * - - (* xO *) - | H : context [ Zpos (xO ?a) ] |- _ => - let isp := isPcst a in - match isp with - | true => change (Zpos (xO a)) with (Zpos' (xO a)) in H - | _ => rewrite (Pos2Z.inj_xO a) in H - end - | |- context [ Zpos (xO ?a) ] => - let isp := isPcst a in - match isp with - | true => change (Zpos (xO a)) with (Zpos' (xO a)) - | _ => rewrite (Pos2Z.inj_xO a) - end - (* xI *) - | H : context [ Zpos (xI ?a) ] |- _ => - let isp := isPcst a in - match isp with - | true => change (Zpos (xI a)) with (Zpos' (xI a)) in H - | _ => rewrite (Pos2Z.inj_xI a) in H - end - | |- context [ Zpos (xI ?a) ] => - let isp := isPcst a in - match isp with - | true => change (Zpos (xI a)) with (Zpos' (xI a)) - | _ => rewrite (Pos2Z.inj_xI a) - end - - (* xI : nothing to do, just prevent adding a useless positivity condition *) - | H : context [ Zpos xH ] |- _ => hide_Zpos xH - | |- context [ Zpos xH ] => hide_Zpos xH - - (* atoms of type positive : we add a positivity condition (if not already there) *) - | _ : 0 < Zpos ?a |- _ => hide_Zpos a - | _ : context [ Zpos ?a ] |- _ => pose proof (Pos2Z.is_pos a); hide_Zpos a - | |- context [ Zpos ?a ] => pose proof (Pos2Z.is_pos a); hide_Zpos a - end. - -#[deprecated( note = "Use 'zify' instead")] -Ltac zify_positive := - repeat zify_positive_rel; repeat zify_positive_op; unfold Zpos',Zneg' in *. - - - - - -(* IV) conversion from N to Z *) - -Definition Z_of_N' := Z.of_N. - -Ltac hide_Z_of_N t := - let z := fresh "z" in set (z:=Z.of_N t) in *; - change Z.of_N with Z_of_N' in z; - unfold z in *; clear z. - -#[deprecated( note = "Use 'zify' instead")] -Ltac zify_N_rel := - match goal with - (* I: equalities *) - | x := ?t : N |- _ => let h := fresh "heq_" x in pose proof (eq_refl : x = t) as h; clearbody x - | |- (@eq N ?a ?b) => apply (N2Z.inj a b) (* shortcut *) - | H : context [ @eq N ?a ?b ] |- _ => rewrite <- (N2Z.inj_iff a b) in H - | |- context [ @eq N ?a ?b ] => rewrite <- (N2Z.inj_iff a b) - (* II: less than *) - | H : context [ (?a < ?b)%N ] |- _ => rewrite (N2Z.inj_lt a b) in H - | |- context [ (?a < ?b)%N ] => rewrite (N2Z.inj_lt a b) - (* III: less or equal *) - | H : context [ (?a <= ?b)%N ] |- _ => rewrite (N2Z.inj_le a b) in H - | |- context [ (?a <= ?b)%N ] => rewrite (N2Z.inj_le a b) - (* IV: greater than *) - | H : context [ (?a > ?b)%N ] |- _ => rewrite (N2Z.inj_gt a b) in H - | |- context [ (?a > ?b)%N ] => rewrite (N2Z.inj_gt a b) - (* V: greater or equal *) - | H : context [ (?a >= ?b)%N ] |- _ => rewrite (N2Z.inj_ge a b) in H - | |- context [ (?a >= ?b)%N ] => rewrite (N2Z.inj_ge a b) - end. - -#[deprecated( note = "Use 'zify' instead")] -Ltac zify_N_op := - match goal with - (* misc type conversions: nat to positive *) - | H : context [ Z.of_N (N.of_nat ?a) ] |- _ => rewrite (nat_N_Z a) in H - | |- context [ Z.of_N (N.of_nat ?a) ] => rewrite (nat_N_Z a) - | H : context [ Z.of_N (Z.abs_N ?a) ] |- _ => rewrite (N2Z.inj_abs_N a) in H - | |- context [ Z.of_N (Z.abs_N ?a) ] => rewrite (N2Z.inj_abs_N a) - | H : context [ Z.of_N (Npos ?a) ] |- _ => rewrite (N2Z.inj_pos a) in H - | |- context [ Z.of_N (Npos ?a) ] => rewrite (N2Z.inj_pos a) - | H : context [ Z.of_N N0 ] |- _ => change (Z.of_N N0) with Z0 in H - | |- context [ Z.of_N N0 ] => change (Z.of_N N0) with Z0 - - (* N.add -> Z.add *) - | H : context [ Z.of_N (N.add ?a ?b) ] |- _ => rewrite (N2Z.inj_add a b) in H - | |- context [ Z.of_N (N.add ?a ?b) ] => rewrite (N2Z.inj_add a b) - - (* N.min -> Z.min *) - | H : context [ Z.of_N (N.min ?a ?b) ] |- _ => rewrite (N2Z.inj_min a b) in H - | |- context [ Z.of_N (N.min ?a ?b) ] => rewrite (N2Z.inj_min a b) - - (* N.max -> Z.max *) - | H : context [ Z.of_N (N.max ?a ?b) ] |- _ => rewrite (N2Z.inj_max a b) in H - | |- context [ Z.of_N (N.max ?a ?b) ] => rewrite (N2Z.inj_max a b) - - (* N.sub -> Z.max 0 (Z.sub ... ...) *) - | H : context [ Z.of_N (N.sub ?a ?b) ] |- _ => rewrite (N2Z.inj_sub_max a b) in H - | |- context [ Z.of_N (N.sub ?a ?b) ] => rewrite (N2Z.inj_sub_max a b) - - (* pred -> minus ... -1 -> Z.max (Z.sub ... -1) 0 *) - | H : context [ Z.of_N (N.pred ?a) ] |- _ => rewrite (N.pred_sub a) in H - | |- context [ Z.of_N (N.pred ?a) ] => rewrite (N.pred_sub a) - - (* N.succ -> Z.succ *) - | H : context [ Z.of_N (N.succ ?a) ] |- _ => rewrite (N2Z.inj_succ a) in H - | |- context [ Z.of_N (N.succ ?a) ] => rewrite (N2Z.inj_succ a) - - (* N.mul -> Z.mul and a positivity hypothesis *) - | H : context [ Z.of_N (N.mul ?a ?b) ] |- _ => - pose proof (N2Z.is_nonneg (N.mul a b)); rewrite (N2Z.inj_mul a b) in * - | |- context [ Z.of_N (N.mul ?a ?b) ] => - pose proof (N2Z.is_nonneg (N.mul a b)); rewrite (N2Z.inj_mul a b) in * - - (* N.div -> Z.div and a positivity hypothesis *) - | H : context [ Z.of_N (N.div ?a ?b) ] |- _ => - pose proof (N2Z.is_nonneg (N.div a b)); rewrite (N2Z.inj_div a b) in * - | |- context [ Z.of_N (N.div ?a ?b) ] => - pose proof (N2Z.is_nonneg (N.div a b)); rewrite (N2Z.inj_div a b) in * - - (* N.modulo -> Z.rem / Z.modulo and a positivity hypothesis (N.modulo agrees with Z.modulo on everything except 0; so we pose both the non-zero proof for this agreement, but also replace things with [Z.rem]) *) - | H : context [ Z.of_N (N.modulo ?a ?b) ] |- _ => - pose proof (N2Z.is_nonneg (N.modulo a b)); - pose proof (@Z.quot_div_nonneg (Z.of_N a) (Z.of_N b) (N2Z.is_nonneg a)); - pose proof (@Z.rem_mod_nonneg (Z.of_N a) (Z.of_N b) (N2Z.is_nonneg a)); - rewrite (N2Z.inj_rem a b) in * - | |- context [ Z.of_N (N.div ?a ?b) ] => - pose proof (N2Z.is_nonneg (N.modulo a b)); - pose proof (@Z.quot_div_nonneg (Z.of_N a) (Z.of_N b) (N2Z.is_nonneg a)); - pose proof (@Z.rem_mod_nonneg (Z.of_N a) (Z.of_N b) (N2Z.is_nonneg a)); - rewrite (N2Z.inj_rem a b) in * - - (* atoms of type N : we add a positivity condition (if not already there) *) - | _ : 0 <= Z.of_N ?a |- _ => hide_Z_of_N a - | _ : context [ Z.of_N ?a ] |- _ => pose proof (N2Z.is_nonneg a); hide_Z_of_N a - | |- context [ Z.of_N ?a ] => pose proof (N2Z.is_nonneg a); hide_Z_of_N a - end. - -#[deprecated( note = "Use 'zify' instead")] -Ltac zify_N := repeat zify_N_rel; repeat zify_N_op; unfold Z_of_N' in *. - -(** The complete Z-ification tactic *) - Require Import ZifyClasses ZifyInst. Require Zify. |
