Skip to content
Snippets Groups Projects
Commit 018e8b64 authored by Jacques-Henri Jourdan's avatar Jacques-Henri Jourdan
Browse files

Some tricks.

parent 07c7a553
No related branches found
No related tags found
No related merge requests found
Pipeline #
......@@ -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.
......@@ -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
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment