From 018e8b64c245432c847e6d6d619963e0fad74448 Mon Sep 17 00:00:00 2001 From: Jacques-Henri Jourdan <jacques-henri.jourdan@normalesup.org> Date: Wed, 14 Dec 2016 10:41:23 +0100 Subject: [PATCH] Some tricks. --- theories/typing/product.v | 24 +++++++++++++++++++++--- theories/typing/uninit.v | 14 ++++++++------ 2 files changed, 29 insertions(+), 9 deletions(-) diff --git a/theories/typing/product.v b/theories/typing/product.v index 5efbb086..e415a891 100644 --- a/theories/typing/product.v +++ b/theories/typing/product.v @@ -118,9 +118,17 @@ Section product. Qed. Definition product := foldr product2 unit. - (* Given that in practice, product will be used with concrete lists, - there should be no need to declare [Copy] and [Proper] instances - for [product]. *) + + Global Instance product_mono E L: + Proper (Forall2 (subtype E L) ==> subtype E L) product. + Proof. intros ??. induction 1. done. by simpl; f_equiv. Qed. + Global Instance product_proper E L: + Proper (Forall2 (eqtype E L) ==> eqtype E L) product. + Proof. intros ??. induction 1. done. by simpl; f_equiv. Qed. + (* FIXME : this instance is never going to be used, because Forall is + not a typeclass. *) + Global Instance product_copy tys : Forall Copy tys → Copy (product tys). + Proof. induction 1; apply _. Qed. End product. Arguments product : simpl never. @@ -174,4 +182,14 @@ Section typing. unfold product. induction tyl1; simpl; last by f_equiv. induction tyl2. by rewrite left_id. by rewrite /= -assoc; f_equiv. Qed. + + Lemma eqtype_prod_nil_flatten E L tyl1 tyl2 : + eqtype E L (Π(Πtyl1 :: tyl2)) (Π(tyl1 ++ tyl2)). + Proof. apply (eqtype_prod_flatten _ _ []). Qed. + Lemma eqtype_prod_flatten_nil E L tyl1 tyl2 : + eqtype E L (Π(tyl1 ++ [Πtyl2])) (Π(tyl1 ++ tyl2)). + Proof. by rewrite (eqtype_prod_flatten E L tyl1 tyl2 []) app_nil_r. Qed. + Lemma eqtype_prod_app E L tyl1 tyl2 : + eqtype E L (Π[Πtyl1; Πtyl2]) (Π(tyl1 ++ tyl2)). + Proof. by rewrite -eqtype_prod_flatten_nil -eqtype_prod_nil_flatten. Qed. End typing. diff --git a/theories/typing/uninit.v b/theories/typing/uninit.v index 7fc9fa0e..2d496c4d 100644 --- a/theories/typing/uninit.v +++ b/theories/typing/uninit.v @@ -11,14 +11,16 @@ Section uninit. Definition uninit (n : nat) : type := Π(replicate n uninit_1). + Global Instance uninit_copy n : Copy (uninit n). + Proof. apply product_copy, Forall_replicate, _. Qed. + + Lemma uninit_sz n : ty_size (uninit n) = n. + Proof. induction n. done. simpl. by f_equal. Qed. + Lemma eqtype_uninit_product E L ns : eqtype E L (uninit (foldr plus 0%nat ns)) (Π(uninit <$> ns)). Proof. - induction ns as [|n ns IH]. done. - rewrite /= /uninit replicate_plus (eqtype_prod_flatten E L []). - induction n. done. rewrite /product /=. by f_equiv. + induction ns as [|n ns IH]. done. revert IH. + by rewrite /= /uninit replicate_plus eqtype_prod_nil_flatten -!eqtype_prod_app=>->. Qed. - - Lemma uninit_sz n : ty_size (uninit n) = n. - Proof. induction n. done. simpl. by f_equal. Qed. End uninit. \ No newline at end of file -- GitLab