diff --git a/theories/typing/examples/diverging_static.v b/theories/typing/examples/diverging_static.v index 0d23942f2d0669555489a864a56b8541ac52c174..fd1ec12b113b9a73cc0f0fda6b822b8de592658d 100644 --- a/theories/typing/examples/diverging_static.v +++ b/theories/typing/examples/diverging_static.v @@ -37,4 +37,39 @@ Section diverging_static. iIntros "!# *". inv_vec args. simpl_subst. iApply type_jump; solve_typing. Qed. + + (** Variant of the above where the lifetime is in an arbitrary invariant + position. This is incompatible with branding! + + Our [TyWf] is not working well for type constructors (we have no way of + representing the fact that well-formedness is somewhat "uniform"), so we + instead work with a constant [Euser] of lifetime inclusion assumptions (in + general it could change with the type parameter but only in monotone ways) + and a single [κuser] of lifetimes that have to be alive (making κuser a + list would require some induction-based reasoning principles that we do + not have, but showing that it works for one lifetime is enough to + demonstrate the principle). *) + Lemma diverging_static_loop_type' (T : lft → type) F κuser Euser call_once + `{!TyWf F} : + let _ : (∀ κ, TyWf (T κ)) := λ κ, {| ty_lfts := [κ; κuser]; ty_wf_E := Euser |} in + (∀ κ1 κ2, (T κ1).(ty_size) = (T κ2).(ty_size)) → + (∀ E L κ1 κ2, lctx_lft_eq E L κ1 κ2 → subtype E L (T κ1) (T κ2)) → + typed_val call_once (fn(∅; F, T static) → unit) → + typed_val (diverging_static_loop call_once) + (fn(∀ α, ∅; T α, F) → ∅). + Proof. + intros HWf HTsz HTsub Hf E L. iApply type_fn; [solve_typing..|]. iIntros "/= !#". + iIntros (α ϝ ret arg). inv_vec arg=>x f. simpl_subst. + iApply type_let; [apply Hf|solve_typing|]. iIntros (call); simpl_subst. + iApply (type_cont [_] [] (λ r, [(r!!!0%fin:val) ◁ box (unit)])). + { iIntros (k). simpl_subst. + iApply type_equivalize_lft_static. + iApply (type_call [ϝ] ()); solve_typing. } + iIntros "!# *". inv_vec args=>r. simpl_subst. + iApply (type_cont [] [] (λ r, [])). + { iIntros (kloop). simpl_subst. iApply type_jump; solve_typing. } + iIntros "!# *". inv_vec args. simpl_subst. + iApply type_jump; solve_typing. + Qed. + End diverging_static. diff --git a/theories/typing/lft_contexts.v b/theories/typing/lft_contexts.v index 6f62322c4a3a20acc1ddfe7fb7126b048cb6374d..14744c72e85a1884766261c1abaff29f1ed1efe1 100644 --- a/theories/typing/lft_contexts.v +++ b/theories/typing/lft_contexts.v @@ -332,6 +332,8 @@ Hint Resolve elctx_sat_nil elctx_sat_lft_incl elctx_sat_app elctx_sat_refl : lrust_typing. +Hint Extern 10 (lctx_lft_eq _ _ _ _) => split : lrust_typing. + Hint Resolve elctx_sat_submseteq | 100 : lrust_typing. Hint Opaque elctx_sat lctx_lft_alive lctx_lft_incl : lrust_typing. diff --git a/theories/typing/shr_bor.v b/theories/typing/shr_bor.v index cf899f62b334f43b7c45a3458569cec94193e526..99690914f786a7ae128999bdddffbac90bb6b349 100644 --- a/theories/typing/shr_bor.v +++ b/theories/typing/shr_bor.v @@ -66,8 +66,8 @@ Section typing. lctx_lft_incl E L κ2 κ1 → subtype E L ty1 ty2 → subtype E L (&shr{κ1}ty1) (&shr{κ2}ty2). Proof. by intros; apply shr_mono. Qed. - Lemma shr_proper' E L κ ty1 ty2 : - eqtype E L ty1 ty2 → eqtype E L (&shr{κ}ty1) (&shr{κ}ty2). + Lemma shr_proper' E L κ1 κ2 ty1 ty2 : + lctx_lft_eq E L κ1 κ2 → eqtype E L ty1 ty2 → eqtype E L (&shr{κ1}ty1) (&shr{κ2}ty2). Proof. by intros; apply shr_proper. Qed. Lemma tctx_reborrow_shr E L p ty κ κ' : diff --git a/theories/typing/uniq_bor.v b/theories/typing/uniq_bor.v index 95d7aaf4182cd90ff1012869d34eae6461d586d9..6ec339ef2ed195b9cf874e698d84b11b47f93d30 100644 --- a/theories/typing/uniq_bor.v +++ b/theories/typing/uniq_bor.v @@ -104,8 +104,8 @@ Section typing. lctx_lft_incl E L κ2 κ1 → eqtype E L ty1 ty2 → subtype E L (&uniq{κ1}ty1) (&uniq{κ2}ty2). Proof. by intros; apply uniq_mono. Qed. - Lemma uniq_proper' E L κ ty1 ty2 : - eqtype E L ty1 ty2 → eqtype E L (&uniq{κ}ty1) (&uniq{κ}ty2). + Lemma uniq_proper' E L κ1 κ2 ty1 ty2 : + lctx_lft_eq E L κ1 κ2 → eqtype E L ty1 ty2 → eqtype E L (&uniq{κ1}ty1) (&uniq{κ2}ty2). Proof. by intros; apply uniq_proper. Qed. Lemma tctx_reborrow_uniq E L p ty κ κ' :