diff --git a/theories/lang/tactics.v b/theories/lang/tactics.v index 21ea93df05520ace50909d0c045e8e5ebef3308e..31301b9da52746d3e62357cf2090e69827e2e865 100644 --- a/theories/lang/tactics.v +++ b/theories/lang/tactics.v @@ -185,7 +185,8 @@ Ltac solve_to_val := match goal with | |- to_val ?e = Some ?v => let e' := W.of_expr e in change (to_val (W.to_expr e') = Some v); - apply W.to_val_Some; simpl; unfold W.to_expr; unlock; reflexivity + apply W.to_val_Some; simpl; unfold W.to_expr; + ((unlock; exact eq_refl) || (idtac; exact eq_refl)) | |- is_Some (to_val ?e) => let e' := W.of_expr e in change (is_Some (to_val (W.to_expr e'))); apply W.to_val_is_Some, (bool_decide_unpack _); vm_compute; exact I diff --git a/theories/typing/base.v b/theories/typing/base.v index d892ae6159b834b9d89f456196a731372c396a5b..a987f68f4bc185484c60b184199258d7706efd26 100644 --- a/theories/typing/base.v +++ b/theories/typing/base.v @@ -14,13 +14,8 @@ Ltac solve_typing := (typeclasses eauto with lrust_typing typeclass_instances core; fail) || (typeclasses eauto with lrust_typing lrust_typing_merge typeclass_instances core; fail). -Lemma of_val_unlock v e : of_val v = e → of_val (locked v) = e. -Proof. by unlock. Qed. - Hint Constructors Forall Forall2 elem_of_list : lrust_typing. -Hint Resolve - of_val_unlock - submseteq_cons submseteq_inserts_l submseteq_inserts_r +Hint Resolve submseteq_cons submseteq_inserts_l submseteq_inserts_r : lrust_typing. Hint Extern 1 (@eq nat _ (Z.to_nat _)) => diff --git a/theories/typing/function.v b/theories/typing/function.v index 87ac86eff00d82f3bec29de2bf036ccadc9e71fc..231257729b9469fd845cbf43ab5335f6a70d5786 100644 --- a/theories/typing/function.v +++ b/theories/typing/function.v @@ -397,10 +397,9 @@ Section typing. Qed. Lemma type_rec {A} E L fb (argsb : list binder) ef e n - (fp : A → fn_params n) T `{!CopyC T, !SendC T} : - ef = (funrec: fb argsb := e)%E → + (fp : A → fn_params n) T `{!CopyC T, !SendC T, Closed _ e} : + IntoVal ef (funrec: fb argsb := e) → n = length argsb → - Closed (fb :b: "return" :b: argsb +b+ []) e → □ (∀ x ϝ (f : val) k (args : vec val (length argsb)), typed_body ((fp x).(fp_E) ϝ) [ϝ ⊑ₗ []] [k ◁cont([ϝ ⊑ₗ []], λ v : vec _ 1, [v!!!0 ◁ box (fp x).(fp_ty)])] @@ -410,8 +409,7 @@ Section typing. (subst_v (fb :: BNamed "return" :: argsb) (f ::: k ::: args) e)) -∗ typed_instruction_ty E L T ef (fn fp). Proof. - iIntros (-> -> Hc) "#Hbody". iIntros (tid) "#LFT _ $ $ #HT". iApply wp_value. - { unfold IntoVal. simpl. rewrite ->(decide_left Hc). done. } + iIntros (<-%of_to_val ->) "#Hbody". iIntros (tid) "#LFT _ $ $ #HT". iApply wp_value. rewrite tctx_interp_singleton. iLöb as "IH". iExists _. iSplit. { simpl. rewrite decide_left. done. } iExists fb, _, argsb, e, _. iSplit. done. iSplit. done. iNext. @@ -422,10 +420,9 @@ Section typing. Qed. Lemma type_fn {A} E L (argsb : list binder) ef e n - (fp : A → fn_params n) T `{!CopyC T, !SendC T} : - ef = (funrec: <> argsb := e)%E → + (fp : A → fn_params n) T `{!CopyC T, !SendC T, Closed _ e} : + IntoVal ef (funrec: <> argsb := e) → n = length argsb → - Closed ("return" :b: argsb +b+ []) e → □ (∀ x ϝ k (args : vec val (length argsb)), typed_body ((fp x).(fp_E) ϝ) [ϝ ⊑ₗ []] [k ◁cont([ϝ ⊑ₗ []], λ v : vec _ 1, [v!!!0 ◁ box (fp x).(fp_ty)])] @@ -434,7 +431,7 @@ Section typing. (subst_v (BNamed "return" :: argsb) (k ::: args) e)) -∗ typed_instruction_ty E L T ef (fn fp). Proof. - iIntros (???) "#He". iApply type_rec; try done. iIntros "!# *". + iIntros (??) "#He". iApply type_rec; try done. iIntros "!# *". iApply typed_body_mono; last iApply "He"; try done. eapply contains_tctx_incl. by constructor. Qed. diff --git a/theories/typing/lib/panic.v b/theories/typing/lib/panic.v index 9050768b6b96d769bf195edbb5d110bd52e70d34..a827be9c9e2eeb0dac3fb26a9958a7e5a95912e3 100644 --- a/theories/typing/lib/panic.v +++ b/theories/typing/lib/panic.v @@ -18,7 +18,7 @@ Section panic. Lemma panic_type : typed_val panic (fn(∅) → emp). Proof. - intros E L. iApply type_fn; [solve_typing..|]. iIntros "!# *". + intros E L. iApply type_fn; [done|]. iIntros "!# *". inv_vec args. iIntros (tid) "LFT HE Hna HL Hk HT /=". simpl_subst. by iApply wp_value. Qed.