From e978da8c41d8a3c19a29036d9c569fbe2a4616b0 Mon Sep 17 00:00:00 2001 From: Samuel Mimram Date: Fri, 16 Jun 2006 14:41:51 +0000 Subject: Imported Upstream version 8.0pl3+8.1beta --- theories/Sorting/PermutEq.v | 241 +++++++++++++++++++++++++++++++++++++++ theories/Sorting/PermutSetoid.v | 243 ++++++++++++++++++++++++++++++++++++++++ theories/Sorting/Permutation.v | 148 +++++++++++++++++++----- 3 files changed, 601 insertions(+), 31 deletions(-) create mode 100644 theories/Sorting/PermutEq.v create mode 100644 theories/Sorting/PermutSetoid.v (limited to 'theories/Sorting') diff --git a/theories/Sorting/PermutEq.v b/theories/Sorting/PermutEq.v new file mode 100644 index 00000000..e56ff27d --- /dev/null +++ b/theories/Sorting/PermutEq.v @@ -0,0 +1,241 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* 0 < multiplicity (list_contents l) a. +Proof. +induction l. +simpl. +split; inversion 1. +simpl. +split; intros. +inversion_clear H. +subst a0. +destruct (eq_dec a a) as [_|H]; auto with arith; destruct H; auto. +destruct (eq_dec a a0) as [H1|H1]; auto with arith; simpl. +rewrite <- IHl; auto. +destruct (eq_dec a a0); auto. +simpl in H. +right; rewrite IHl; auto. +Qed. + +Lemma multiplicity_In_O : + forall l a, ~ In a l -> multiplicity (list_contents l) a = 0. +Proof. +intros l a; rewrite multiplicity_In; + destruct (multiplicity (list_contents l) a); auto. +destruct 1; auto with arith. +Qed. + +Lemma multiplicity_In_S : + forall l a, In a l -> multiplicity (list_contents l) a >= 1. +Proof. +intros l a; rewrite multiplicity_In; auto. +Qed. + +Lemma multiplicity_NoDup : + forall l, NoDup l <-> (forall a, multiplicity (list_contents l) a <= 1). +Proof. +induction l. +simpl. +split; auto with arith. +intros; apply NoDup_nil. +split; simpl. +inversion_clear 1. +rewrite IHl in H1. +intros; destruct (eq_dec a a0) as [H2|H2]; simpl; auto. +subst a0. +rewrite multiplicity_In_O; auto. +intros; constructor. +rewrite multiplicity_In. +generalize (H a). +destruct (eq_dec a a) as [H0|H0]. +destruct (multiplicity (list_contents l) a); auto with arith. +simpl; inversion 1. +inversion H3. +destruct H0; auto. +rewrite IHl; intros. +generalize (H a0); auto with arith. +destruct (eq_dec a a0); simpl; auto with arith. +Qed. + +Lemma NoDup_permut : + forall l l', NoDup l -> NoDup l' -> + (forall x, In x l <-> In x l') -> permutation l l'. +Proof. +intros. +red; unfold meq; intros. +rewrite multiplicity_NoDup in H, H0. +generalize (H a) (H0 a) (H1 a); clear H H0 H1. +do 2 rewrite multiplicity_In. +destruct 3; omega. +Qed. + +(** Permutation is compatible with In. *) +Lemma permut_In_In : + forall l1 l2 e, permutation l1 l2 -> In e l1 -> In e l2. +Proof. +unfold Permutation.permutation, meq; intros l1 l2 e P IN. +generalize (P e); clear P. +destruct (In_dec eq_dec e l2) as [H|H]; auto. +rewrite (multiplicity_In_O _ _ H). +intros. +generalize (multiplicity_In_S _ _ IN). +rewrite H0. +inversion 1. +Qed. + +Lemma permut_cons_In : + forall l1 l2 e, permutation (e :: l1) l2 -> In e l2. +Proof. +intros; eapply permut_In_In; eauto. +red; auto. +Qed. + +(** Permutation of an empty list. *) +Lemma permut_nil : + forall l, permutation l nil -> l = nil. +Proof. +intro l; destruct l as [ | e l ]; trivial. +assert (In e (e::l)) by (red; auto). +intro Abs; generalize (permut_In_In _ Abs H). +inversion 1. +Qed. + +(** When used with [eq], this permutation notion is equivalent to + the one defined in [List.v]. *) + +Lemma permutation_Permutation : + forall l l', Permutation l l' <-> permutation l l'. +Proof. +split. +induction 1. +apply permut_refl. +apply permut_cons; auto. +change (permutation (y::x::l) ((x::nil)++y::l)). +apply permut_add_cons_inside; simpl; apply permut_refl. +apply permut_tran with l'; auto. +revert l'. +induction l. +intros. +rewrite (permut_nil (permut_sym H)). +apply Permutation_refl. +intros. +destruct (In_split _ _ (permut_cons_In H)) as (h2,(t2,H1)). +subst l'. +apply Permutation_cons_app. +apply IHl. +apply permut_remove_hd with a; auto. +Qed. + +(** Permutation for short lists. *) + +Lemma permut_length_1: + forall a b, permutation (a :: nil) (b :: nil) -> a=b. +Proof. +intros a b; unfold Permutation.permutation, meq; intro P; +generalize (P b); clear P; simpl. +destruct (eq_dec b b) as [H|H]; [ | destruct H; auto]. +destruct (eq_dec a b); simpl; auto; intros; discriminate. +Qed. + +Lemma permut_length_2 : + forall a1 b1 a2 b2, permutation (a1 :: b1 :: nil) (a2 :: b2 :: nil) -> + (a1=a2) /\ (b1=b2) \/ (a1=b2) /\ (a2=b1). +Proof. +intros a1 b1 a2 b2 P. +assert (H:=permut_cons_In P). +inversion_clear H. +left; split; auto. +apply permut_length_1. +red; red; intros. +generalize (P a); clear P; simpl. +destruct (eq_dec a1 a) as [H2|H2]; + destruct (eq_dec a2 a) as [H3|H3]; auto. +destruct H3; transitivity a1; auto. +destruct H2; transitivity a2; auto. +right. +inversion_clear H0; [|inversion H]. +split; auto. +apply permut_length_1. +red; red; intros. +generalize (P a); clear P; simpl. +destruct (eq_dec a1 a) as [H2|H2]; + destruct (eq_dec b2 a) as [H3|H3]; auto. +simpl; rewrite <- plus_n_Sm; inversion 1; auto. +destruct H3; transitivity a1; auto. +destruct H2; transitivity b2; auto. +Qed. + +(** Permutation is compatible with length. *) +Lemma permut_length : + forall l1 l2, permutation l1 l2 -> length l1 = length l2. +Proof. +induction l1; intros l2 H. +rewrite (permut_nil (permut_sym H)); auto. +destruct (In_split _ _ (permut_cons_In H)) as (h2,(t2,H1)). +subst l2. +rewrite app_length. +simpl; rewrite <- plus_n_Sm; f_equal. +rewrite <- app_length. +apply IHl1. +apply permut_remove_hd with a; auto. +Qed. + +Variable B : Set. +Variable eqB_dec : forall x y:B, { x=y }+{ ~x=y }. + +(** Permutation is compatible with map. *) + +Lemma permutation_map : + forall f l1 l2, permutation l1 l2 -> + Permutation.permutation _ eqB_dec (map f l1) (map f l2). +Proof. +intros f; induction l1. +intros l2 P; rewrite (permut_nil (permut_sym P)); apply permut_refl. +intros l2 P. +simpl. +destruct (In_split _ _ (permut_cons_In P)) as (h2,(t2,H1)). +subst l2. +rewrite map_app. +simpl. +apply permut_add_cons_inside. +rewrite <- map_app. +apply IHl1; auto. +apply permut_remove_hd with a; auto. +Qed. + +End Perm. + diff --git a/theories/Sorting/PermutSetoid.v b/theories/Sorting/PermutSetoid.v new file mode 100644 index 00000000..46ea088f --- /dev/null +++ b/theories/Sorting/PermutSetoid.v @@ -0,0 +1,243 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* eqA y x. +Variable eqA_trans : forall x y z, eqA x y -> eqA y z -> eqA x z. + +(** we can use [multiplicity] to define [InA] and [NoDupA]. *) + +Lemma multiplicity_InA : + forall l a, InA eqA a l <-> 0 < multiplicity (list_contents l) a. +Proof. +induction l. +simpl. +split; inversion 1. +simpl. +split; intros. +inversion_clear H. +destruct (eqA_dec a a0) as [_|H1]; auto with arith. +destruct H1; auto. +destruct (eqA_dec a a0); auto with arith. +simpl; rewrite <- IHl; auto. +destruct (eqA_dec a a0) as [H0|H0]; auto. +simpl in H. +constructor 2; rewrite IHl; auto. +Qed. + +Lemma multiplicity_InA_O : + forall l a, ~ InA eqA a l -> multiplicity (list_contents l) a = 0. +Proof. +intros l a; rewrite multiplicity_InA; +destruct (multiplicity (list_contents l) a); auto with arith. +destruct 1; auto with arith. +Qed. + +Lemma multiplicity_InA_S : + forall l a, InA eqA a l -> multiplicity (list_contents l) a >= 1. +Proof. +intros l a; rewrite multiplicity_InA; auto with arith. +Qed. + +Lemma multiplicity_NoDupA : forall l, + NoDupA eqA l <-> (forall a, multiplicity (list_contents l) a <= 1). +Proof. +induction l. +simpl. +split; auto with arith. +split; simpl. +inversion_clear 1. +rewrite IHl in H1. +intros; destruct (eqA_dec a a0) as [H2|H2]; simpl; auto. +rewrite multiplicity_InA_O; auto. +swap H0. +apply InA_eqA with a0; auto. +intros; constructor. +rewrite multiplicity_InA. +generalize (H a). +destruct (eqA_dec a a) as [H0|H0]. +destruct (multiplicity (list_contents l) a); auto with arith. +simpl; inversion 1. +inversion H3. +destruct H0; auto. +rewrite IHl; intros. +generalize (H a0); auto with arith. +destruct (eqA_dec a a0); simpl; auto with arith. +Qed. + + +(** Permutation is compatible with InA. *) +Lemma permut_InA_InA : + forall l1 l2 e, permutation l1 l2 -> InA eqA e l1 -> InA eqA e l2. +Proof. +intros l1 l2 e. +do 2 rewrite multiplicity_InA. +unfold Permutation.permutation, meq. +intros H;rewrite H; auto. +Qed. + +Lemma permut_cons_InA : + forall l1 l2 e, permutation (e :: l1) l2 -> InA eqA e l2. +Proof. +intros; apply (permut_InA_InA (e:=e) H); auto. +Qed. + +(** Permutation of an empty list. *) +Lemma permut_nil : + forall l, permutation l nil -> l = nil. +Proof. +intro l; destruct l as [ | e l ]; trivial. +assert (InA eqA e (e::l)) by auto. +intro Abs; generalize (permut_InA_InA Abs H). +inversion 1. +Qed. + +(** Permutation for short lists. *) + +Lemma permut_length_1: + forall a b, permutation (a :: nil) (b :: nil) -> eqA a b. +Proof. +intros a b; unfold Permutation.permutation, meq; intro P; +generalize (P b); clear P; simpl. +destruct (eqA_dec b b) as [H|H]; [ | destruct H; auto]. +destruct (eqA_dec a b); simpl; auto; intros; discriminate. +Qed. + +Lemma permut_length_2 : + forall a1 b1 a2 b2, permutation (a1 :: b1 :: nil) (a2 :: b2 :: nil) -> + (eqA a1 a2) /\ (eqA b1 b2) \/ (eqA a1 b2) /\ (eqA a2 b1). +Proof. +intros a1 b1 a2 b2 P. +assert (H:=permut_cons_InA P). +inversion_clear H. +left; split; auto. +apply permut_length_1. +red; red; intros. +generalize (P a); clear P; simpl. +destruct (eqA_dec a1 a) as [H2|H2]; + destruct (eqA_dec a2 a) as [H3|H3]; auto. +destruct H3; apply eqA_trans with a1; auto. +destruct H2; apply eqA_trans with a2; auto. +right. +inversion_clear H0; [|inversion H]. +split; auto. +apply permut_length_1. +red; red; intros. +generalize (P a); clear P; simpl. +destruct (eqA_dec a1 a) as [H2|H2]; + destruct (eqA_dec b2 a) as [H3|H3]; auto. +simpl; rewrite <- plus_n_Sm; inversion 1; auto. +destruct H3; apply eqA_trans with a1; auto. +destruct H2; apply eqA_trans with b2; auto. +Qed. + +(** Permutation is compatible with length. *) +Lemma permut_length : + forall l1 l2, permutation l1 l2 -> length l1 = length l2. +Proof. +induction l1; intros l2 H. +rewrite (permut_nil (permut_sym H)); auto. +assert (H0:=permut_cons_InA H). +destruct (InA_split H0) as (h2,(b,(t2,(H1,H2)))). +subst l2. +rewrite app_length. +simpl; rewrite <- plus_n_Sm; f_equal. +rewrite <- app_length. +apply IHl1. +apply permut_remove_hd with b. +apply permut_tran with (a::l1); auto. +revert H1; unfold Permutation.permutation, meq; simpl. +intros; f_equal; auto. +destruct (eqA_dec b a0) as [H2|H2]; + destruct (eqA_dec a a0) as [H3|H3]; auto. +destruct H3; apply eqA_trans with b; auto. +destruct H2; apply eqA_trans with a; auto. +Qed. + +Lemma NoDupA_eqlistA_permut : + forall l l', NoDupA eqA l -> NoDupA eqA l' -> + eqlistA eqA l l' -> permutation l l'. +Proof. +intros. +red; unfold meq; intros. +rewrite multiplicity_NoDupA in H, H0. +generalize (H a) (H0 a) (H1 a); clear H H0 H1. +do 2 rewrite multiplicity_InA. +destruct 3; omega. +Qed. + + +Variable B : Set. +Variable eqB : B->B->Prop. +Variable eqB_dec : forall x y:B, { eqB x y }+{ ~eqB x y }. +Variable eqB_trans : forall x y z, eqB x y -> eqB y z -> eqB x z. + +(** Permutation is compatible with map. *) + +Lemma permut_map : + forall f, + (forall x y, eqA x y -> eqB (f x) (f y)) -> + forall l1 l2, permutation l1 l2 -> + Permutation.permutation _ eqB_dec (map f l1) (map f l2). +Proof. +intros f; induction l1. +intros l2 P; rewrite (permut_nil (permut_sym P)); apply permut_refl. +intros l2 P. +simpl. +assert (H0:=permut_cons_InA P). +destruct (InA_split H0) as (h2,(b,(t2,(H1,H2)))). +subst l2. +rewrite map_app. +simpl. +apply permut_tran with (f b :: map f l1). +revert H1; unfold Permutation.permutation, meq; simpl. +intros; f_equal; auto. +destruct (eqB_dec (f b) a0) as [H2|H2]; + destruct (eqB_dec (f a) a0) as [H3|H3]; auto. +destruct H3; apply eqB_trans with (f b); auto. +destruct H2; apply eqB_trans with (f a); auto. +apply permut_add_cons_inside. +rewrite <- map_app. +apply IHl1; auto. +apply permut_remove_hd with b. +apply permut_tran with (a::l1); auto. +revert H1; unfold Permutation.permutation, meq; simpl. +intros; f_equal; auto. +destruct (eqA_dec b a0) as [H2|H2]; + destruct (eqA_dec a a0) as [H3|H3]; auto. +destruct H3; apply eqA_trans with b; auto. +destruct H2; apply eqA_trans with a; auto. +Qed. + +End Perm. diff --git a/theories/Sorting/Permutation.v b/theories/Sorting/Permutation.v index b3287cd1..0f2e02b5 100644 --- a/theories/Sorting/Permutation.v +++ b/theories/Sorting/Permutation.v @@ -6,30 +6,39 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(*i $Id: Permutation.v 5920 2004-07-16 20:01:26Z herbelin $ i*) +(*i $Id: Permutation.v 8823 2006-05-16 16:17:43Z letouzey $ i*) Require Import Relations. Require Import List. Require Import Multiset. +Require Import Arith. + +(** This file define a notion of permutation for lists, based on multisets: + there exists a permutation between two lists iff every elements have + the same multiplicities in the two lists. + + Unlike [List.Permutation], the present notion of permutation requires + a decidable equality. At the same time, this definition can be used + with a non-standard equality, whereas [List.Permutation] cannot. + + The present file contains basic results, obtained without any particular + assumption on the decidable equality used. + + File [PermutSetoid] contains additional results about permutations + with respect to an setoid equality (i.e. an equivalence relation). + + Finally, file [PermutEq] concerns Coq equality : this file is similar + to the previous one, but proves in addition that [List.Permutation] + and [permutation] are equivalent in this context. +*) Set Implicit Arguments. Section defs. Variable A : Set. -Variable leA : relation A. Variable eqA : relation A. - -Let gtA (x y:A) := ~ leA x y. - -Hypothesis leA_dec : forall x y:A, {leA x y} + {~ leA x y}. Hypothesis eqA_dec : forall x y:A, {eqA x y} + {~ eqA x y}. -Hypothesis leA_refl : forall x y:A, eqA x y -> leA x y. -Hypothesis leA_trans : forall x y z:A, leA x y -> leA y z -> leA x z. -Hypothesis leA_antisym : forall x y:A, leA x y -> leA y x -> eqA x y. - -Hint Resolve leA_refl: default. -Hint Immediate eqA_dec leA_dec leA_antisym: default. Let emptyBag := EmptyBag A. Let singletonBag := SingletonBag _ eqA_dec. @@ -63,6 +72,12 @@ unfold permutation in |- *; auto with datatypes. Qed. Hint Resolve permut_refl. +Lemma permut_sym : + forall l1 l2 : list A, permutation l1 l2 -> permutation l2 l1. +Proof. +unfold permutation, meq; intros; apply sym_eq; trivial. +Qed. + Lemma permut_tran : forall l m n:list A, permutation l m -> permutation m n -> permutation l n. Proof. @@ -70,51 +85,122 @@ unfold permutation in |- *; intros. apply meq_trans with (list_contents m); auto with datatypes. Qed. -Lemma permut_right : +Lemma permut_cons : forall l m:list A, permutation l m -> forall a:A, permutation (a :: l) (a :: m). Proof. unfold permutation in |- *; simpl in |- *; auto with datatypes. Qed. -Hint Resolve permut_right. +Hint Resolve permut_cons. Lemma permut_app : forall l l' m m':list A, permutation l l' -> permutation m m' -> permutation (l ++ m) (l' ++ m'). Proof. unfold permutation in |- *; intros. -apply meq_trans with (munion (list_contents l) (list_contents m)); +apply meq_trans with (munion (list_contents l) (list_contents m)); auto with datatypes. -apply meq_trans with (munion (list_contents l') (list_contents m')); +apply meq_trans with (munion (list_contents l') (list_contents m')); auto with datatypes. apply meq_trans with (munion (list_contents l') (list_contents m)); auto with datatypes. Qed. Hint Resolve permut_app. -Lemma permut_cons : - forall l m:list A, - permutation l m -> forall a:A, permutation (a :: l) (a :: m). +Lemma permut_add_inside : + forall a l1 l2 l3 l4, + permutation (l1 ++ l2) (l3 ++ l4) -> + permutation (l1 ++ a :: l2) (l3 ++ a :: l4). Proof. -intros l m H a. -change (permutation ((a :: nil) ++ l) ((a :: nil) ++ m)) in |- *. -apply permut_app; auto with datatypes. +unfold permutation, meq in *; intros. +generalize (H a0); clear H. +do 4 rewrite list_contents_app. +simpl. +destruct (eqA_dec a a0); simpl; auto with arith. +do 2 rewrite <- plus_n_Sm; f_equal; auto. +Qed. + +Lemma permut_add_cons_inside : + forall a l l1 l2, + permutation l (l1 ++ l2) -> + permutation (a :: l) (l1 ++ a :: l2). +Proof. +intros; +replace (a :: l) with (nil ++ a :: l); trivial; +apply permut_add_inside; trivial. Qed. -Hint Resolve permut_cons. Lemma permut_middle : forall (l m:list A) (a:A), permutation (a :: l ++ m) (l ++ a :: m). Proof. -unfold permutation in |- *. -simple induction l; simpl in |- *; auto with datatypes. -intros. -apply meq_trans with - (munion (singletonBag a) - (munion (singletonBag a0) (list_contents (l0 ++ m)))); - auto with datatypes. -apply munion_perm_left; auto with datatypes. +intros; apply permut_add_cons_inside; auto. Qed. Hint Resolve permut_middle. +Lemma permut_sym_app : + forall l1 l2, permutation (l1 ++ l2) (l2 ++ l1). +Proof. +intros l1 l2; +unfold permutation, meq; +intro a; do 2 rewrite list_contents_app; simpl; +auto with arith. +Qed. + +Lemma permut_rev : + forall l, permutation l (rev l). +Proof. +induction l. +simpl; auto. +simpl. +apply permut_add_cons_inside. +rewrite <- app_nil_end; auto. +Qed. + +(** Some inversion results. *) +Lemma permut_conv_inv : + forall e l1 l2, permutation (e :: l1) (e :: l2) -> permutation l1 l2. +Proof. +intros e l1 l2; unfold permutation, meq; simpl; intros H a; +generalize (H a); apply plus_reg_l. +Qed. + +Lemma permut_app_inv1 : + forall l l1 l2, permutation (l1 ++ l) (l2 ++ l) -> permutation l1 l2. +Proof. +intros l l1 l2; unfold permutation, meq; simpl; +intros H a; generalize (H a); clear H. +do 2 rewrite list_contents_app. +simpl. +intros; apply plus_reg_l with (multiplicity (list_contents l) a). +rewrite plus_comm; rewrite H; rewrite plus_comm. +trivial. +Qed. + +Lemma permut_app_inv2 : + forall l l1 l2, permutation (l ++ l1) (l ++ l2) -> permutation l1 l2. +Proof. +intros l l1 l2; unfold permutation, meq; simpl; +intros H a; generalize (H a); clear H. +do 2 rewrite list_contents_app. +simpl. +intros; apply plus_reg_l with (multiplicity (list_contents l) a). +trivial. +Qed. + +Lemma permut_remove_hd : + forall l l1 l2 a, + permutation (a :: l) (l1 ++ a :: l2) -> permutation l (l1 ++ l2). +Proof. +intros l l1 l2 a; unfold permutation, meq; simpl; intros H a0; generalize (H a0); clear H. +do 2 rewrite list_contents_app; simpl; intro H. +apply plus_reg_l with (if eqA_dec a a0 then 1 else 0). +rewrite H; clear H. +symmetry; rewrite plus_comm. +repeat rewrite <- plus_assoc; f_equal. +apply plus_comm. +Qed. + End defs. +(* For compatibilty *) +Notation permut_right := permut_cons. Unset Implicit Arguments. -- cgit v1.2.3