Skip to content
Snippets Groups Projects
Commit d3414ad2 authored by Ralf Jung's avatar Ralf Jung
Browse files

rename some eqtype lemmas

parent 2fc79e0d
No related branches found
No related tags found
No related merge requests found
Pipeline #
......@@ -176,20 +176,21 @@ Section typing.
- simpl. iFrame.
Qed.
Lemma eqtype_prod_flatten E L tyl1 tyl2 tyl3 :
Lemma prod_flatten E L tyl1 tyl2 tyl3 :
eqtype E L (Π(tyl1 ++ Π tyl2 :: tyl3)) (Π(tyl1 ++ tyl2 ++ tyl3)).
Proof.
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 :
Lemma prod_flatten_l 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 :
Proof. apply (prod_flatten _ _ []). Qed.
Lemma prod_flatten_r 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 :
Proof. by rewrite (prod_flatten E L tyl1 tyl2 []) app_nil_r. Qed.
Lemma 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.
Proof. by rewrite -prod_flatten_r -prod_flatten_l. Qed.
End typing.
......@@ -15,7 +15,7 @@ Section shr_bor.
Qed.
Global Instance subtype_shr_bor_mono E L :
Proper (lctx_lft_incl E L --> subtype E L ==> subtype E L) shr_bor.
Proper (flip (lctx_lft_incl E L) ==> subtype E L ==> subtype E L) shr_bor.
Proof.
intros κ1 κ2 ty1 ty2 Hty. apply subtype_simple_type.
iIntros (??) "#LFT #HE #HL H". iDestruct ( with "HE HL") as "#Hκ".
......@@ -25,7 +25,7 @@ Section shr_bor.
by iApply "Hs1".
Qed.
Global Instance subtype_shr_bor_mono' E L :
Proper (lctx_lft_incl E L ==> subtype E L --> flip (subtype E L)) shr_bor.
Proper (lctx_lft_incl E L ==> flip (subtype E L) ==> flip (subtype E L)) shr_bor.
Proof. intros ??????. by apply subtype_shr_bor_mono. Qed.
Global Instance subtype_shr_bor_proper E L κ :
Proper (eqtype E L ==> eqtype E L) (shr_bor κ).
......
......@@ -17,10 +17,10 @@ Section uninit.
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 :
Lemma uninit_product E L ns :
eqtype E L (uninit (foldr plus 0%nat ns)) (Π(uninit <$> ns)).
Proof.
induction ns as [|n ns IH]. done. revert IH.
by rewrite /= /uninit replicate_plus eqtype_prod_nil_flatten -!eqtype_prod_app=>->.
by rewrite /= /uninit replicate_plus prod_flatten_l -!prod_app=>->.
Qed.
End uninit.
\ No newline at end of file
End uninit.
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