diff options
| author | letouzey | 2010-02-09 17:44:52 +0000 |
|---|---|---|
| committer | letouzey | 2010-02-09 17:44:52 +0000 |
| commit | bf90d39cec401f5daad2eb26c915ceba65e1a5cc (patch) | |
| tree | ee1eb266a33d1f9d16af1268136f60ee72c5bc04 | |
| parent | 38dac30c6877122634e7b34ec7cd1b6ab2b67ebb (diff) | |
NPeano improved, subsumes NatOrderedType
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@12717 85f007b7-540e-0410-9357-904b9bb8a0f7
| -rw-r--r-- | theories/Arith/Max.v | 33 | ||||
| -rw-r--r-- | theories/Arith/Min.v | 33 | ||||
| -rw-r--r-- | theories/Arith/MinMax.v | 69 | ||||
| -rw-r--r-- | theories/Arith/NatOrderedType.v | 64 | ||||
| -rw-r--r-- | theories/Arith/vo.itarget | 1 | ||||
| -rw-r--r-- | theories/Numbers/Natural/Peano/NPeano.v | 152 | ||||
| -rw-r--r-- | theories/Structures/OrdersEx.v | 4 |
7 files changed, 141 insertions, 215 deletions
diff --git a/theories/Arith/Max.v b/theories/Arith/Max.v index 3d7fe9fc2d..e49251a717 100644 --- a/theories/Arith/Max.v +++ b/theories/Arith/Max.v @@ -8,34 +8,35 @@ (*i $Id$ i*) -(** THIS FILE IS DEPRECATED. Use [MinMax] instead. *) +(** THIS FILE IS DEPRECATED. Use [NPeano] and [MinMax] instead. *) +Require Import NPeano. Require Export MinMax. Local Open Scope nat_scope. Implicit Types m n p : nat. -Notation max := MinMax.max (only parsing). +Notation max := NPeano.max (only parsing). Definition max_0_l := max_0_l. Definition max_0_r := max_0_r. Definition succ_max_distr := succ_max_distr. Definition plus_max_distr_l := plus_max_distr_l. Definition plus_max_distr_r := plus_max_distr_r. -Definition max_case_strong := max_case_strong. -Definition max_spec := max_spec. -Definition max_dec := max_dec. -Definition max_case := max_case. -Definition max_idempotent := max_id. -Definition max_assoc := max_assoc. -Definition max_comm := max_comm. -Definition max_l := max_l. -Definition max_r := max_r. -Definition le_max_l := le_max_l. -Definition le_max_r := le_max_r. -Definition max_lub_l := max_lub_l. -Definition max_lub_r := max_lub_r. -Definition max_lub := max_lub. +Definition max_case_strong := Nat.max_case_strong. +Definition max_spec := Nat.max_spec. +Definition max_dec := Nat.max_dec. +Definition max_case := Nat.max_case. +Definition max_idempotent := Nat.max_id. +Definition max_assoc := Nat.max_assoc. +Definition max_comm := Nat.max_comm. +Definition max_l := Nat.max_l. +Definition max_r := Nat.max_r. +Definition le_max_l := Nat.le_max_l. +Definition le_max_r := Nat.le_max_r. +Definition max_lub_l := Nat.max_lub_l. +Definition max_lub_r := Nat.max_lub_r. +Definition max_lub := Nat.max_lub. (* begin hide *) (* Compatibility *) diff --git a/theories/Arith/Min.v b/theories/Arith/Min.v index c52fc0dd08..44060b56c1 100644 --- a/theories/Arith/Min.v +++ b/theories/Arith/Min.v @@ -8,34 +8,35 @@ (*i $Id$ i*) -(** THIS FILE IS DEPRECATED. Use [MinMax] instead. *) +(** THIS FILE IS DEPRECATED. Use [NPeano] and [MinMax] instead. *) +Require Import NPeano. Require Export MinMax. Open Local Scope nat_scope. Implicit Types m n p : nat. -Notation min := MinMax.min (only parsing). +Notation min := NPeano.min (only parsing). Definition min_0_l := min_0_l. Definition min_0_r := min_0_r. Definition succ_min_distr := succ_min_distr. Definition plus_min_distr_l := plus_min_distr_l. Definition plus_min_distr_r := plus_min_distr_r. -Definition min_case_strong := min_case_strong. -Definition min_spec := min_spec. -Definition min_dec := min_dec. -Definition min_case := min_case. -Definition min_idempotent := min_id. -Definition min_assoc := min_assoc. -Definition min_comm := min_comm. -Definition min_l := min_l. -Definition min_r := min_r. -Definition le_min_l := le_min_l. -Definition le_min_r := le_min_r. -Definition min_glb_l := min_glb_l. -Definition min_glb_r := min_glb_r. -Definition min_glb := min_glb. +Definition min_case_strong := Nat.min_case_strong. +Definition min_spec := Nat.min_spec. +Definition min_dec := Nat.min_dec. +Definition min_case := Nat.min_case. +Definition min_idempotent := Nat.min_id. +Definition min_assoc := Nat.min_assoc. +Definition min_comm := Nat.min_comm. +Definition min_l := Nat.min_l. +Definition min_r := Nat.min_r. +Definition le_min_l := Nat.le_min_l. +Definition le_min_r := Nat.le_min_r. +Definition min_glb_l := Nat.min_glb_l. +Definition min_glb_r := Nat.min_glb_r. +Definition min_glb := Nat.min_glb. (* begin hide *) (* Compatibility *) diff --git a/theories/Arith/MinMax.v b/theories/Arith/MinMax.v index 6e86a88c59..9ce43ab8ce 100644 --- a/theories/Arith/MinMax.v +++ b/theories/Arith/MinMax.v @@ -6,61 +6,10 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -Require Import Orders NatOrderedType GenericMinMax. - -(** * Maximum and Minimum of two natural numbers *) - -Fixpoint max n m : nat := - match n, m with - | O, _ => m - | S n', O => n - | S n', S m' => S (max n' m') - end. - -Fixpoint min n m : nat := - match n, m with - | O, _ => 0 - | S n', O => 0 - | S n', S m' => S (min n' m') - end. - -(** These functions implement indeed a maximum and a minimum *) - -Lemma max_l : forall x y, y<=x -> max x y = x. -Proof. - induction x; destruct y; simpl; auto with arith. -Qed. - -Lemma max_r : forall x y, x<=y -> max x y = y. -Proof. - induction x; destruct y; simpl; auto with arith. -Qed. - -Lemma min_l : forall x y, x<=y -> min x y = x. -Proof. - induction x; destruct y; simpl; auto with arith. -Qed. - -Lemma min_r : forall x y, y<=x -> min x y = y. -Proof. - induction x; destruct y; simpl; auto with arith. -Qed. - - -Module NatHasMinMax <: HasMinMax Nat_as_OT. - Definition max := max. - Definition min := min. - Definition max_l := max_l. - Definition max_r := max_r. - Definition min_l := min_l. - Definition min_r := min_r. -End NatHasMinMax. - -(** We obtain hence all the generic properties of [max] and [min], - see file [GenericMinMax] or use SearchAbout. *) - -Module Export MMP := UsualMinMaxProperties Nat_as_OT NatHasMinMax. +Require Import Orders NPeano. +(** The Definition of maximum and minimum of two natural numbers + is now in NPeano, as well as generic properties. *) (** * Properties specific to the [nat] domain *) @@ -88,26 +37,26 @@ Proof. auto. Qed. Lemma plus_max_distr_l : forall n m p, max (p + n) (p + m) = p + max n m. Proof. -intros. apply max_monotone. repeat red; auto with arith. +intros. apply Nat.max_monotone. repeat red; auto with arith. Qed. Lemma plus_max_distr_r : forall n m p, max (n + p) (m + p) = max n m + p. Proof. -intros. apply max_monotone with (f:=fun x => x + p). +intros. apply Nat.max_monotone with (f:=fun x => x + p). repeat red; auto with arith. Qed. Lemma plus_min_distr_l : forall n m p, min (p + n) (p + m) = p + min n m. Proof. -intros. apply min_monotone. repeat red; auto with arith. +intros. apply Nat.min_monotone. repeat red; auto with arith. Qed. Lemma plus_min_distr_r : forall n m p, min (n + p) (m + p) = min n m + p. Proof. -intros. apply min_monotone with (f:=fun x => x + p). +intros. apply Nat.min_monotone with (f:=fun x => x + p). repeat red; auto with arith. Qed. Hint Resolve - max_l max_r le_max_l le_max_r - min_l min_r le_min_l le_min_r : arith v62. + Nat.max_l Nat.max_r Nat.le_max_l Nat.le_max_r + Nat.min_l Nat.min_r Nat.le_min_l Nat.le_min_r : arith v62. diff --git a/theories/Arith/NatOrderedType.v b/theories/Arith/NatOrderedType.v deleted file mode 100644 index df5b37e019..0000000000 --- a/theories/Arith/NatOrderedType.v +++ /dev/null @@ -1,64 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) -(* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(************************************************************************) - -Require Import Lt Peano_dec Compare_dec EqNat - Equalities Orders OrdersTac. - - -(** * DecidableType structure for Peano numbers *) - -Module Nat_as_UBE <: UsualBoolEq. - Definition t := nat. - Definition eq := @eq nat. - Definition eqb := beq_nat. - Definition eqb_eq := beq_nat_true_iff. -End Nat_as_UBE. - -Module Nat_as_DT <: UsualDecidableTypeFull := Make_UDTF Nat_as_UBE. - -(** Note that the last module fulfills by subtyping many other - interfaces, such as [DecidableType] or [EqualityType]. *) - - - -(** * OrderedType structure for Peano numbers *) - -Module Nat_as_OT <: OrderedTypeFull. - Include Nat_as_DT. - Definition lt := lt. - Definition le := le. - Definition compare := nat_compare. - - Instance lt_strorder : StrictOrder lt. - Proof. split; [ exact lt_irrefl | exact lt_trans ]. Qed. - - Instance lt_compat : Proper (Logic.eq==>Logic.eq==>iff) lt. - Proof. repeat red; intros; subst; auto. Qed. - - Definition le_lteq := le_lt_or_eq_iff. - Definition compare_spec := nat_compare_spec. - -End Nat_as_OT. - -(** Note that [Nat_as_OT] can also be seen as a [UsualOrderedType] - and a [OrderedType] (and also as a [DecidableType]). *) - - - -(** * An [order] tactic for Peano numbers *) - -Module NatOrder := OTF_to_OrderTac Nat_as_OT. -Ltac nat_order := NatOrder.order. - -(** Note that [nat_order] is domain-agnostic: it will not prove - [1<=2] or [x<=x+x], but rather things like [x<=y -> y<=x -> x=y]. *) - -Section Test. -Let test : forall x y : nat, x<=y -> y<=x -> x=y. -Proof. nat_order. Qed. -End Test. diff --git a/theories/Arith/vo.itarget b/theories/Arith/vo.itarget index c3f29d2143..d8bff0ab91 100644 --- a/theories/Arith/vo.itarget +++ b/theories/Arith/vo.itarget @@ -19,5 +19,4 @@ Mult.vo Peano_dec.vo Plus.vo Wf_nat.vo -NatOrderedType.vo MinMax.vo diff --git a/theories/Numbers/Natural/Peano/NPeano.v b/theories/Numbers/Natural/Peano/NPeano.v index d42bb810d2..e5c6824193 100644 --- a/theories/Numbers/Natural/Peano/NPeano.v +++ b/theories/Numbers/Natural/Peano/NPeano.v @@ -10,11 +10,44 @@ (*i $Id$ i*) -Require Import Arith MinMax NAxioms NProperties. +Require Import Peano Peano_dec Compare_dec EqNat NAxioms NProperties NDiv. + +(** Some functions not already encountered: min max div mod *) + +Fixpoint max n m : nat := + match n, m with + | O, _ => m + | S n', O => n + | S n', S m' => S (max n' m') + end. + +Fixpoint min n m : nat := + match n, m with + | O, _ => 0 + | S n', O => 0 + | S n', S m' => S (min n' m') + end. + +Definition divF div x y := if leb y x then S (div (x-y) y) else 0. +Definition modF mod x y := if leb y x then mod (x-y) y else x. +Definition initF (_ _ : nat) := 0. + +Fixpoint loop {A} (F:A->A)(i:A) (n:nat) : A := + match n with + | 0 => i + | S n => F (loop F i n) + end. + +Definition div x y := loop divF initF x x y. +Definition modulo x y := loop modF initF x x y. +Infix "/" := div : nat_scope. +Infix "mod" := modulo (at level 40, no associativity) : nat_scope. + (** * Implementation of [NAxiomsSig] by [nat] *) -Module NPeanoAxiomsMod <: NAxiomsSig. +Module Nat + <: NAxiomsSig <: UsualDecidableTypeFull <: OrderedTypeFull <: TotalOrder. (** Bi-directional induction. *) @@ -57,7 +90,7 @@ Qed. Theorem sub_succ_r : forall n m : nat, n - (S m) = pred (n - m). Proof. -intros n m; induction n m using nat_double_ind; simpl; auto. apply sub_0_r. +induction n; destruct m; simpl; auto. apply sub_0_r. Qed. Theorem mul_0_l : forall n : nat, 0 * n = 0. @@ -67,49 +100,61 @@ Qed. Theorem mul_succ_l : forall n m : nat, S n * m = n * m + m. Proof. -intros n m; now rewrite plus_comm. +assert (add_S_r : forall n m, n+S m = S(n+m)) by (induction n; auto). +assert (add_comm : forall n m, n+m = m+n). + induction n; simpl; auto. intros; rewrite add_S_r; auto. +intros n m; now rewrite add_comm. Qed. (** Order on natural numbers *) Program Instance lt_wd : Proper (eq==>eq==>iff) lt. -Theorem lt_eq_cases : forall n m : nat, n <= m <-> n < m \/ n = m. +Theorem lt_succ_r : forall n m : nat, n < S m <-> n <= m. Proof. -intros n m; split. -apply le_lt_or_eq. -intro H; destruct H as [H | H]. -now apply lt_le_weak. rewrite H; apply le_refl. +unfold lt. +split; [|induction 1; auto]. +assert (le_pred : forall n m, n <= m -> pred n <= pred m). + induction 1 as [|m' H IH]; auto. + destruct m'. inversion H; subst; auto. simpl; auto. +exact (le_pred (S n) (S m)). Qed. -Theorem lt_irrefl : forall n : nat, ~ (n < n). + +Theorem lt_eq_cases : forall n m : nat, n <= m <-> n < m \/ n = m. Proof. -exact lt_irrefl. +split. +inversion 1; auto. rewrite lt_succ_r; auto. +destruct 1; [|subst; auto]. rewrite <- lt_succ_r; auto. Qed. -Theorem lt_succ_r : forall n m : nat, n < S m <-> n <= m. +Theorem lt_irrefl : forall n : nat, ~ (n < n). Proof. -intros n m; split; [apply lt_n_Sm_le | apply le_lt_n_Sm]. +induction n. intro H; inversion H. rewrite lt_succ_r; auto. Qed. Theorem min_l : forall n m : nat, n <= m -> min n m = n. Proof. -exact min_l. +induction n; destruct m; simpl; auto. inversion 1. +rewrite (lt_succ_r n m). auto. Qed. Theorem min_r : forall n m : nat, m <= n -> min n m = m. Proof. -exact min_r. +induction n; destruct m; simpl; auto. inversion 1. +rewrite (lt_succ_r m n). auto. Qed. Theorem max_l : forall n m : nat, m <= n -> max n m = n. Proof. -exact max_l. +induction n; destruct m; simpl; auto. inversion 1. +rewrite (lt_succ_r m n). auto. Qed. Theorem max_r : forall n m : nat, n <= m -> max n m = m. Proof. -exact max_r. +induction n; destruct m; simpl; auto. inversion 1. +rewrite (lt_succ_r n m). auto. Qed. (** Facts specific to natural numbers, not integers. *) @@ -149,6 +194,8 @@ Qed. Definition t := nat. Definition eq := @eq nat. +Definition eqb := beq_nat. +Definition compare := nat_compare. Definition zero := 0. Definition succ := S. Definition pred := pred. @@ -160,30 +207,16 @@ Definition le := le. Definition min := min. Definition max := max. -End NPeanoAxiomsMod. - -(** Now we apply the largest property functor *) +Definition eqb_eq := beq_nat_true_iff. +Definition compare_spec := nat_compare_spec. +Definition eq_dec := eq_nat_dec. -Module Export NPeanoPropMod := NPropFunct NPeanoAxiomsMod. +(** Generic Properties *) +Include NPropFunct + <+ UsualMinMaxLogicalProperties <+ UsualMinMaxDecProperties. - -(** Euclidean Division *) - -Definition divF div x y := if leb y x then S (div (x-y) y) else 0. -Definition modF mod x y := if leb y x then mod (x-y) y else x. -Definition initF (_ _ : nat) := 0. - -Fixpoint loop {A} (F:A->A)(i:A) (n:nat) : A := - match n with - | 0 => i - | S n => F (loop F i n) - end. - -Definition div x y := loop divF initF x x y. -Definition modulo x y := loop modF initF x x y. -Infix "/" := div : nat_scope. -Infix "mod" := modulo (at level 40, no associativity) : nat_scope. +(** Proofs of specification for [div] and [mod]. *) Lemma div_mod : forall x y, y<>0 -> x = y*(x/y) + x mod y. Proof. @@ -195,14 +228,13 @@ Proof. simpl; unfold divF at 1, modF at 1. intros. destruct (leb y x) as [ ]_eqn:L; - [apply leb_complete in L | apply leb_complete_conv in L]. + [apply leb_complete in L | apply leb_complete_conv in L; now nzsimpl]. rewrite mul_succ_r, <- add_assoc, (add_comm y), add_assoc. rewrite <- IHn; auto. symmetry; apply sub_add; auto. - rewrite <- NPeanoAxiomsMod.lt_succ_r. + rewrite <- lt_succ_r. apply lt_le_trans with x; auto. - apply lt_minus; auto. rewrite <- neq_0_lt_0; auto. - nzsimpl; auto. + apply sub_lt; auto. rewrite <- neq_0_lt_0; auto. Qed. Lemma mod_upper_bound : forall x y, y<>0 -> x mod y < y. @@ -216,22 +248,30 @@ Proof. destruct (leb y x) as [ ]_eqn:L; [apply leb_complete in L | apply leb_complete_conv in L]; auto. apply IHn; auto. - rewrite <- NPeanoAxiomsMod.lt_succ_r. + rewrite <- lt_succ_r. apply lt_le_trans with x; auto. - apply lt_minus; auto. rewrite <- neq_0_lt_0; auto. + apply sub_lt; auto. rewrite <- neq_0_lt_0; auto. Qed. -Require Import NDiv. +Definition div := div. +Definition modulo := modulo. +Program Instance div_wd : Proper (eq==>eq==>eq) div. +Program Instance mod_wd : Proper (eq==>eq==>eq) modulo. + +(** Generic properties of [div] and [mod] *) + +Include NDivPropFunct. + +End Nat. + +(** [Nat] contains an [order] tactic for natural numbers *) -Module NDivMod <: NDivSig. - Include NPeanoAxiomsMod. - Definition div := div. - Definition modulo := modulo. - Definition div_mod := div_mod. - Definition mod_upper_bound := mod_upper_bound. - Local Obligation Tactic := simpl_relation. - Program Instance div_wd : Proper (eq==>eq==>eq) div. - Program Instance mod_wd : Proper (eq==>eq==>eq) modulo. -End NDivMod. +(** Note that [Nat.order] is domain-agnostic: it will not prove + [1<=2] or [x<=x+x], but rather things like [x<=y -> y<=x -> x=y]. *) -Module Export NDivPropMod := NDivPropFunct NDivMod NPeanoPropMod. +Section TestOrder. + Let test : forall x y, x<=y -> y<=x -> x=y. + Proof. + Nat.order. + Qed. +End TestOrder. diff --git a/theories/Structures/OrdersEx.v b/theories/Structures/OrdersEx.v index 90840ce1f6..b0702f6818 100644 --- a/theories/Structures/OrdersEx.v +++ b/theories/Structures/OrdersEx.v @@ -13,7 +13,7 @@ (* $Id$ *) -Require Import Orders NatOrderedType POrderedType NArith +Require Import Orders NPeano POrderedType NArith ZArith RelationPairs EqualitiesFacts. (** * Examples of Ordered Type structures. *) @@ -21,7 +21,7 @@ Require Import Orders NatOrderedType POrderedType NArith (** Ordered Type for [nat], [Positive], [N], [Z] with the usual order. *) -Module Nat_as_OT := NatOrderedType.Nat_as_OT. +Module Nat_as_OT := NPeano.Nat. Module Positive_as_OT := POrderedType.Positive_as_OT. Module N_as_OT := NArith.N. Module Z_as_OT := ZArith.Z. |
