From 4f0ad99adb04e7f2888e75f2a10e8c916dde179b Mon Sep 17 00:00:00 2001 From: letouzey Date: Tue, 3 Nov 2009 08:24:06 +0000 Subject: OrderedType implementation for various numerical datatypes + min/max structures - A richer OrderedTypeFull interface : OrderedType + predicate "le" - Implementations {Nat,N,P,Z,Q}OrderedType.v, also providing "order" tactics - By the way: as suggested by S. Lescuyer, specification of compare is now inductive - GenericMinMax: axiomatisation + properties of min and max out of OrderedTypeFull structures. - MinMax.v, {Z,P,N,Q}minmax.v are specialization of GenericMinMax, with also some domain-specific results, and compatibility layer with already existing results. - Some ML code of plugins had to be adapted, otherwise wrong "eq", "lt" or simimlar constants were found by functions like coq_constant. - Beware of the aliasing problems: for instance eq:=@eq t instead of eq:=@eq M.t in Make_UDT made (r)omega stopped working (Z_as_OT.t instead of Z in statement of Zmax_spec). - Some Morphism declaration are now ambiguous: switch to new syntax anyway. - Misc adaptations of FSets/MSets - Classes/RelationPairs.v: from two relations over A and B, we inspect relations over A*B and their properties in terms of classes. git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@12461 85f007b7-540e-0410-9357-904b9bb8a0f7 --- theories/NArith/BinNat.v | 8 ++ theories/NArith/BinPos.v | 7 ++ theories/NArith/NOrderedType.v | 72 +++++++++++++ theories/NArith/Nminmax.v | 226 +++++++++++++++++++++++++++++++++++++++++ theories/NArith/POrderedType.v | 73 +++++++++++++ theories/NArith/Pminmax.v | 225 ++++++++++++++++++++++++++++++++++++++++ 6 files changed, 611 insertions(+) create mode 100644 theories/NArith/NOrderedType.v create mode 100644 theories/NArith/Nminmax.v create mode 100644 theories/NArith/POrderedType.v create mode 100644 theories/NArith/Pminmax.v (limited to 'theories/NArith') diff --git a/theories/NArith/BinNat.v b/theories/NArith/BinNat.v index e02f2817c..813956332 100644 --- a/theories/NArith/BinNat.v +++ b/theories/NArith/BinNat.v @@ -140,6 +140,14 @@ Definition Nmax (n n' : N) := match Ncompare n n' with | Gt => n end. +(** Decidability of equality. *) + +Definition N_eq_dec : forall n m : N, { n = m } + { n <> m }. +Proof. + decide equality. + apply positive_eq_dec. +Defined. + (** convenient induction principles *) Lemma N_ind_double : diff --git a/theories/NArith/BinPos.v b/theories/NArith/BinPos.v index 0a1cd9ab4..20cfb43a3 100644 --- a/theories/NArith/BinPos.v +++ b/theories/NArith/BinPos.v @@ -751,6 +751,13 @@ Proof. destruct (Pcompare_not_Eq p q) as (_,H'); elim H'; auto. Qed. +Lemma Pcompare_eq_iff : forall p q:positive, (p ?= q) Eq = Eq <-> p = q. +Proof. + split. + apply Pcompare_Eq_eq. + intros; subst; apply Pcompare_refl. +Qed. + Lemma Pcompare_Gt_Lt : forall p q:positive, (p ?= q) Gt = Lt -> (p ?= q) Eq = Lt. Proof. diff --git a/theories/NArith/NOrderedType.v b/theories/NArith/NOrderedType.v new file mode 100644 index 000000000..8f69cd656 --- /dev/null +++ b/theories/NArith/NOrderedType.v @@ -0,0 +1,72 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Logic.eq==>iff) Nlt. + Proof. repeat red; intros; subst; auto. Qed. + + Lemma le_lteq : forall x y, x <= y <-> x < y \/ x=y. + Proof. + unfold Nle, Nlt; intros. rewrite <- Ncompare_eq_correct. + destruct (x ?= y); intuition; discriminate. + Qed. + + Lemma compare_spec : forall x y, Cmp eq lt x y (Ncompare x y). + Proof. + intros. + destruct (Ncompare x y) as [ ]_eqn; constructor; auto. + apply Ncompare_Eq_eq; auto. + apply Ngt_Nlt; auto. + Qed. + +End N_as_OT. + +(* Note that [N_as_OT] can also be seen as a [UsualOrderedType] + and a [OrderedType] (and also as a [DecidableType]). *) + + + +(** * An [order] tactic for [N] numbers *) + +Module NOrder := OTF_to_OrderTac N_as_OT. +Ltac n_order := + change (@eq N) with NOrder.OrderElts.eq in *; + NOrder.order. + +(** Note that [n_order] is domain-agnostic: it will not prove + [1<=2] or [x<=x+x], but rather things like [x<=y -> y<=x -> x=y]. *) + diff --git a/theories/NArith/Nminmax.v b/theories/NArith/Nminmax.v new file mode 100644 index 000000000..86e6b609f --- /dev/null +++ b/theories/NArith/Nminmax.v @@ -0,0 +1,226 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Type), + (m<=n -> P n) -> (n<=m -> P m) -> P (Nmax n m). +Proof. intros; apply max_case_strong; auto. congruence. Defined. + +Lemma Nmax_case : forall n m (P:N -> Type), + P n -> P m -> P (Nmax n m). +Proof. intros. apply Nmax_case_strong; auto. Defined. + +Lemma Nmax_monotone: forall f, + (Proper (Nle ==> Nle) f) -> + forall x y, Nmax (f x) (f y) = f (Nmax x y). +Proof. intros; apply max_monotone; auto. congruence. Qed. + +Lemma Nmin_case_strong : forall n m (P:N -> Type), + (m<=n -> P m) -> (n<=m -> P n) -> P (Nmin n m). +Proof. intros; apply min_case_strong; auto. congruence. Defined. + +Lemma Nmin_case : forall n m (P:N -> Type), + P n -> P m -> P (Nmin n m). +Proof. intros. apply Nmin_case_strong; auto. Defined. + +Lemma Nmin_monotone: forall f, + (Proper (Nle ==> Nle) f) -> + forall x y, Nmin (f x) (f y) = f (Nmin x y). +Proof. intros; apply min_monotone; auto. congruence. Qed. + +Lemma Nmax_min_antimonotone : forall f, + Proper (Nle==>Nge) f -> + forall x y, Nmax (f x) (f y) == f (Nmin x y). +Proof. + intros f H. apply max_min_antimonotone. congruence. + intros x x' Hx. red. specialize (H _ _ Hx). clear Hx. + unfold Nle, Nge in *; contradict H. rewrite <- Ncompare_antisym, H; auto. +Qed. + +Lemma Nmin_max_antimonotone : forall f, + Proper (Nle==>Nge) f -> + forall x y, Nmin (f x) (f y) == f (Nmax x y). +Proof. + intros f H. apply min_max_antimonotone. congruence. + intros z z' Hz; red. specialize (H _ _ Hz). clear Hz. + unfold Nle, Nge in *. contradict H. rewrite <- Ncompare_antisym, H; auto. +Qed. + +(** For the other generic properties, we make aliases, + since otherwise SearchAbout misses some of them + (bad interaction with an Include). + See GenericMinMax (or SearchAbout) for the statements. *) + +Definition Nmax_spec_le := max_spec_le. +Definition Nmax_dec := max_dec. +Definition Nmax_unicity := max_unicity. +Definition Nmax_unicity_ext := max_unicity_ext. +Definition Nmax_id := max_id. +Notation Nmax_idempotent := Nmax_id (only parsing). +Definition Nmax_assoc := max_assoc. +Definition Nmax_comm := max_comm. +Definition Nmax_l := max_l. +Definition Nmax_r := max_r. +Definition Nle_max_l := le_max_l. +Definition Nle_max_r := le_max_r. +Definition Nmax_le := max_le. +Definition Nmax_le_iff := max_le_iff. +Definition Nmax_lt_iff := max_lt_iff. +Definition Nmax_lub_l := max_lub_l. +Definition Nmax_lub_r := max_lub_r. +Definition Nmax_lub := max_lub. +Definition Nmax_lub_iff := max_lub_iff. +Definition Nmax_lub_lt := max_lub_lt. +Definition Nmax_lub_lt_iff := max_lub_lt_iff. +Definition Nmax_le_compat_l := max_le_compat_l. +Definition Nmax_le_compat_r := max_le_compat_r. +Definition Nmax_le_compat := max_le_compat. + +Definition Nmin_spec_le := min_spec_le. +Definition Nmin_dec := min_dec. +Definition Nmin_unicity := min_unicity. +Definition Nmin_unicity_ext := min_unicity_ext. +Definition Nmin_id := min_id. +Notation Nmin_idempotent := Nmin_id (only parsing). +Definition Nmin_assoc := min_assoc. +Definition Nmin_comm := min_comm. +Definition Nmin_l := min_l. +Definition Nmin_r := min_r. +Definition Nle_min_l := le_min_l. +Definition Nle_min_r := le_min_r. +Definition Nmin_le := min_le. +Definition Nmin_le_iff := min_le_iff. +Definition Nmin_lt_iff := min_lt_iff. +Definition Nmin_glb_l := min_glb_l. +Definition Nmin_glb_r := min_glb_r. +Definition Nmin_glb := min_glb. +Definition Nmin_glb_iff := min_glb_iff. +Definition Nmin_glb_lt := min_glb_lt. +Definition Nmin_glb_lt_iff := min_glb_lt_iff. +Definition Nmin_le_compat_l := min_le_compat_l. +Definition Nmin_le_compat_r := min_le_compat_r. +Definition Nmin_le_compat := min_le_compat. + +Definition Nmin_max_absorption := min_max_absorption. +Definition Nmax_min_absorption := max_min_absorption. +Definition Nmax_min_distr := max_min_distr. +Definition Nmin_max_distr := min_max_distr. +Definition Nmax_min_modular := max_min_modular. +Definition Nmin_max_modular := min_max_modular. +Definition Nmax_min_disassoc := max_min_disassoc. + + + +(** * Properties specific to the [positive] domain *) + +(** Simplifications *) + +Lemma Nmax_0_l : forall n, Nmax 0 n = n. +Proof. + intros. unfold Nmax. rewrite <- Ncompare_antisym. generalize (Ncompare_0 n). + destruct (n ?= 0); intuition. +Qed. + +Lemma Nmax_0_r : forall n, Nmax n 0 = n. +Proof. intros. rewrite max_comm. apply Nmax_0_l. Qed. + +Lemma Nmin_0_l : forall n, Nmin 0 n = 0. +Proof. + intros. unfold Nmin. rewrite <- Ncompare_antisym. generalize (Ncompare_0 n). + destruct (n ?= 0); intuition. +Qed. + +Lemma Nmin_0_r : forall n, Nmin n 0 = 0. +Proof. intros. rewrite min_comm. apply Nmin_0_l. Qed. + +(** Compatibilities (consequences of monotonicity) *) + +Lemma Nsucc_max_distr : + forall n m, Nsucc (Nmax n m) = Nmax (Nsucc n) (Nsucc m). +Proof. + intros. symmetry. apply Nmax_monotone. + intros x x'. unfold Nle. + rewrite 2 nat_of_Ncompare, 2 nat_of_Nsucc. + simpl; auto. +Qed. + +Lemma Nsucc_min_distr : forall n m, Nsucc (Nmin n m) = Nmin (Nsucc n) (Nsucc m). +Proof. + intros. symmetry. apply Nmin_monotone. + intros x x'. unfold Nle. + rewrite 2 nat_of_Ncompare, 2 nat_of_Nsucc. + simpl; auto. +Qed. + +Lemma Nplus_max_distr_l : forall n m p, Nmax (p + n) (p + m) = p + Nmax n m. +Proof. + intros. apply Nmax_monotone. + intros x x'. unfold Nle. + rewrite 2 nat_of_Ncompare, 2 nat_of_Nplus. + rewrite <- 2 Compare_dec.nat_compare_le. auto with arith. +Qed. + +Lemma Nplus_max_distr_r : forall n m p, Nmax (n + p) (m + p) = Nmax n m + p. +Proof. + intros. rewrite (Nplus_comm n p), (Nplus_comm m p), (Nplus_comm _ p). + apply Nplus_max_distr_l. +Qed. + +Lemma Nplus_min_distr_l : forall n m p, Nmin (p + n) (p + m) = p + Nmin n m. +Proof. + intros. apply Nmin_monotone. + intros x x'. unfold Nle. + rewrite 2 nat_of_Ncompare, 2 nat_of_Nplus. + rewrite <- 2 Compare_dec.nat_compare_le. auto with arith. +Qed. + +Lemma Nplus_min_distr_r : forall n m p, Nmin (n + p) (m + p) = Nmin n m + p. +Proof. + intros. rewrite (Nplus_comm n p), (Nplus_comm m p), (Nplus_comm _ p). + apply Nplus_min_distr_l. +Qed. diff --git a/theories/NArith/POrderedType.v b/theories/NArith/POrderedType.v new file mode 100644 index 000000000..2f89b0e68 --- /dev/null +++ b/theories/NArith/POrderedType.v @@ -0,0 +1,73 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Logic.eq==>iff) Plt. + Proof. repeat red; intros; subst; auto. Qed. + + Lemma le_lteq : forall x y, x <= y <-> x < y \/ x=y. + Proof. + unfold Ple, Plt. intros. + rewrite <- Pcompare_eq_iff. + destruct (Pcompare x y Eq); intuition; discriminate. + Qed. + + Lemma compare_spec : forall x y, Cmp eq lt x y (compare x y). + Proof. + intros; unfold compare. + destruct (Pcompare x y Eq) as [ ]_eqn; constructor. + apply Pcompare_Eq_eq; auto. + auto. + apply ZC1; auto. + Qed. + +End Positive_as_OT. + +(* Note that [Positive_as_OT] can also be seen as a [UsualOrderedType] + and a [OrderedType] (and also as a [DecidableType]). *) + + + +(** * An [order] tactic for positive numbers *) + +Module PositiveOrder := OTF_to_OrderTac Positive_as_OT. +Ltac p_order := + change (@eq positive) with PositiveOrder.OrderElts.eq in *; + PositiveOrder.order. + +(** Note that [p_order] is domain-agnostic: it will not prove + [1<=2] or [x<=x+x], but rather things like [x<=y -> y<=x -> x=y]. *) diff --git a/theories/NArith/Pminmax.v b/theories/NArith/Pminmax.v new file mode 100644 index 000000000..18008d18d --- /dev/null +++ b/theories/NArith/Pminmax.v @@ -0,0 +1,225 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Type), + (m<=n -> P n) -> (n<=m -> P m) -> P (Pmax n m). +Proof. intros; apply max_case_strong; auto. congruence. Defined. + +Lemma Pmax_case : forall n m (P:positive -> Type), + P n -> P m -> P (Pmax n m). +Proof. intros. apply Pmax_case_strong; auto. Defined. + +Lemma Pmax_monotone: forall f, + (Proper (Ple ==> Ple) f) -> + forall x y, Pmax (f x) (f y) = f (Pmax x y). +Proof. intros; apply max_monotone; auto. congruence. Qed. + +Lemma Pmin_case_strong : forall n m (P:positive -> Type), + (m<=n -> P m) -> (n<=m -> P n) -> P (Pmin n m). +Proof. intros; apply min_case_strong; auto. congruence. Defined. + +Lemma Pmin_case : forall n m (P:positive -> Type), + P n -> P m -> P (Pmin n m). +Proof. intros. apply Pmin_case_strong; auto. Defined. + +Lemma Pmin_monotone: forall f, + (Proper (Ple ==> Ple) f) -> + forall x y, Pmin (f x) (f y) = f (Pmin x y). +Proof. intros; apply min_monotone; auto. congruence. Qed. + +Lemma Pmax_min_antimonotone : forall f, + Proper (Ple==>Pge) f -> + forall x y, Pmax (f x) (f y) == f (Pmin x y). +Proof. + intros f H. apply max_min_antimonotone. congruence. + intros z z' Hz; red. specialize (H _ _ Hz). clear Hz. + unfold Ple, Pge in *. contradict H. rewrite ZC4, H; auto. +Qed. + +Lemma Pmin_max_antimonotone : forall f, + Proper (Ple==>Pge) f -> + forall x y, Pmin (f x) (f y) == f (Pmax x y). +Proof. + intros f H. apply min_max_antimonotone. congruence. + intros z z' Hz; red. specialize (H _ _ Hz). clear Hz. + unfold Ple, Pge in *. contradict H. rewrite ZC4, H; auto. +Qed. + +(** For the other generic properties, we make aliases, + since otherwise SearchAbout misses some of them + (bad interaction with an Include). + See GenericMinMax (or SearchAbout) for the statements. *) + +Definition Pmax_spec_le := max_spec_le. +Definition Pmax_dec := max_dec. +Definition Pmax_unicity := max_unicity. +Definition Pmax_unicity_ext := max_unicity_ext. +Definition Pmax_id := max_id. +Notation Pmax_idempotent := Pmax_id (only parsing). +Definition Pmax_assoc := max_assoc. +Definition Pmax_comm := max_comm. +Definition Pmax_l := max_l. +Definition Pmax_r := max_r. +Definition Ple_max_l := le_max_l. +Definition Ple_max_r := le_max_r. +Definition Pmax_le := max_le. +Definition Pmax_le_iff := max_le_iff. +Definition Pmax_lt_iff := max_lt_iff. +Definition Pmax_lub_l := max_lub_l. +Definition Pmax_lub_r := max_lub_r. +Definition Pmax_lub := max_lub. +Definition Pmax_lub_iff := max_lub_iff. +Definition Pmax_lub_lt := max_lub_lt. +Definition Pmax_lub_lt_iff := max_lub_lt_iff. +Definition Pmax_le_compat_l := max_le_compat_l. +Definition Pmax_le_compat_r := max_le_compat_r. +Definition Pmax_le_compat := max_le_compat. + +Definition Pmin_spec_le := min_spec_le. +Definition Pmin_dec := min_dec. +Definition Pmin_unicity := min_unicity. +Definition Pmin_unicity_ext := min_unicity_ext. +Definition Pmin_id := min_id. +Notation Pmin_idempotent := Pmin_id (only parsing). +Definition Pmin_assoc := min_assoc. +Definition Pmin_comm := min_comm. +Definition Pmin_l := min_l. +Definition Pmin_r := min_r. +Definition Ple_min_l := le_min_l. +Definition Ple_min_r := le_min_r. +Definition Pmin_le := min_le. +Definition Pmin_le_iff := min_le_iff. +Definition Pmin_lt_iff := min_lt_iff. +Definition Pmin_glb_l := min_glb_l. +Definition Pmin_glb_r := min_glb_r. +Definition Pmin_glb := min_glb. +Definition Pmin_glb_iff := min_glb_iff. +Definition Pmin_glb_lt := min_glb_lt. +Definition Pmin_glb_lt_iff := min_glb_lt_iff. +Definition Pmin_le_compat_l := min_le_compat_l. +Definition Pmin_le_compat_r := min_le_compat_r. +Definition Pmin_le_compat := min_le_compat. + +Definition Pmin_max_absorption := min_max_absorption. +Definition Pmax_min_absorption := max_min_absorption. +Definition Pmax_min_distr := max_min_distr. +Definition Pmin_max_distr := min_max_distr. +Definition Pmax_min_modular := max_min_modular. +Definition Pmin_max_modular := min_max_modular. +Definition Pmax_min_disassoc := max_min_disassoc. + + +(** * Properties specific to the [positive] domain *) + +(** Simplifications *) + +Lemma Pmax_1_l : forall n, Pmax 1 n = n. +Proof. + intros. unfold Pmax. rewrite ZC4. generalize (Pcompare_1 n). + destruct (n ?= 1); intuition. +Qed. + +Lemma Pmax_1_r : forall n, Pmax n 1 = n. +Proof. intros. rewrite max_comm. apply Pmax_1_l. Qed. + +Lemma Pmin_1_l : forall n, Pmin 1 n = 1. +Proof. + intros. unfold Pmin. rewrite ZC4. generalize (Pcompare_1 n). + destruct (n ?= 1); intuition. +Qed. + +Lemma Pmin_1_r : forall n, Pmin n 1 = 1. +Proof. intros. rewrite min_comm. apply Pmin_1_l. Qed. + +(** Compatibilities (consequences of monotonicity) *) + +Lemma Psucc_max_distr : + forall n m, Psucc (Pmax n m) = Pmax (Psucc n) (Psucc m). +Proof. + intros. symmetry. apply Pmax_monotone. + intros x x'. unfold Ple. + rewrite 2 nat_of_P_compare_morphism, 2 nat_of_P_succ_morphism. + simpl; auto. +Qed. + +Lemma Psucc_min_distr : forall n m, Psucc (Pmin n m) = Pmin (Psucc n) (Psucc m). +Proof. + intros. symmetry. apply Pmin_monotone. + intros x x'. unfold Ple. + rewrite 2 nat_of_P_compare_morphism, 2 nat_of_P_succ_morphism. + simpl; auto. +Qed. + +Lemma Pplus_max_distr_l : forall n m p, Pmax (p + n) (p + m) = p + Pmax n m. +Proof. + intros. apply Pmax_monotone. + intros x x'. unfold Ple. + rewrite 2 nat_of_P_compare_morphism, 2 nat_of_P_plus_morphism. + rewrite <- 2 Compare_dec.nat_compare_le. auto with arith. +Qed. + +Lemma Pplus_max_distr_r : forall n m p, Pmax (n + p) (m + p) = Pmax n m + p. +Proof. + intros. rewrite (Pplus_comm n p), (Pplus_comm m p), (Pplus_comm _ p). + apply Pplus_max_distr_l. +Qed. + +Lemma Pplus_min_distr_l : forall n m p, Pmin (p + n) (p + m) = p + Pmin n m. +Proof. + intros. apply Pmin_monotone. + intros x x'. unfold Ple. + rewrite 2 nat_of_P_compare_morphism, 2 nat_of_P_plus_morphism. + rewrite <- 2 Compare_dec.nat_compare_le. auto with arith. +Qed. + +Lemma Pplus_min_distr_r : forall n m p, Pmin (n + p) (m + p) = Pmin n m + p. +Proof. + intros. rewrite (Pplus_comm n p), (Pplus_comm m p), (Pplus_comm _ p). + apply Pplus_min_distr_l. +Qed. -- cgit v1.2.3