RunST.logrelNN_ctx_closure
From RunST Require Import typing lang rules regions reduction logrel_shared
logrelNN fundamentalNN contextual_refinement.
Ltac fold_interp :=
match goal with
| |- context [ interp_exprNN (interp_arrow (interpNN ?A) (interpNN ?A'))
?B (?C, ?D) ] =>
change (interp_exprNN (interp_arrow (interpNN A) (interpNN A')) B (C, D)) with
(interp_exprNN (interpNN (TArrow A A')) B (C, D))
| |- context [ interp_exprNN (interp_prod (interpNN ?A) (interpNN ?A'))
?B (?C, ?D) ] =>
change (interp_exprNN (interp_prod (interpNN A) (interpNN A')) B (C, D)) with
(interp_exprNN (interpNN (TProd A A')) B (C, D))
| |- context [ interp_exprNN (interp_prod (interpNN ?A) (interpNN ?A'))
?B (?C, ?D) ] =>
change (interp_exprNN (interp_prod (interpNN A) (interpNN A')) B (C, D)) with
(interp_exprNN (interpNN (TProd A A')) B (C, D))
| |- context [ interp_exprNN (interp_sum (interpNN ?A) (interpNN ?A'))
?B (?C, ?D) ] =>
change (interp_exprNN (interp_sum (interpNN A) (interpNN A')) B (C, D)) with
(interp_exprNN (interpNN (TSum A A')) B (C, D))
| |- context [ interp_exprNN (interp_rec (interpNN ?A)) ?B (?C, ?D) ] =>
change (interp_exprNN (interp_rec (interpNN A)) B (C, D)) with
(interp_exprNN (interpNN (TRec A)) B (C, D))
| |- context [ interp_exprNN (interp_forall (interpNN ?A))
?B (?C, ?D) ] =>
change (interp_exprNN (interp_forall (interpNN A)) B (C, D)) with
(interp_exprNN (interpNN (TForall A)) B (C, D))
| |- context [ interp_exprNN (interp_ref ?E (interpNN ?A))
?B (?C, ?D) ] =>
change (interp_exprNN (interp_ref E (interpNN A)) B (C, D)) with
(interp_exprNN (interpNN (TSTref E A)) B (C, D))
end.
Section bin_log_related_under_typed_ctx.
Context `{heapG Σ, LogRelG Σ}.
Lemma bin_log_related_under_typed_ctx Γ e e' τ Γ' τ' K :
(∀ f, e.[upn (length Γ) f] = e) →
(∀ f, e'.[upn (length Γ) f] = e') →
typed_ctx K Γ τ Γ' τ' →
Γ ⊨ e ≤logNN≤ e' : τ → Γ' ⊨ fill_ctx K e ≤logNN≤ fill_ctx K e' : τ'.
Proof.
revert Γ τ Γ' τ' e e'.
induction K as [|k K] => Γ τ Γ' τ' e e' H1 H2; simpl.
- inversion_clear 1; trivial.
- inversion_clear 1 as [|? ? ? ? ? ? ? ? Hx1 Hx2]. intros H3.
specialize (IHK _ _ _ _ e e' H1 H2 Hx2 H3).
inversion Hx1; subst; simpl; try fold_interp.
+ eapply bin_log_related_rec; eauto;
match goal with
H : _ |- _ => eapply (typed_ctx_n_closed _ _ _ _ _ _ _ H)
end.
+ eapply bin_log_related_app; eauto using binary_fundamental.
+ eapply bin_log_related_app; eauto using binary_fundamental.
+ eapply bin_log_related_pair; eauto using binary_fundamental.
+ eapply bin_log_related_pair; eauto using binary_fundamental.
+ eapply bin_log_related_fst; eauto.
+ eapply bin_log_related_snd; eauto.
+ eapply bin_log_related_injl; eauto.
+ eapply bin_log_related_injr; eauto.
+ match goal with
H : typed_ctx_item _ _ _ _ _ |- _ => inversion H; subst
end.
eapply bin_log_related_case;
eauto using binary_fundamental;
match goal with
H : _ |- _ => eapply (typed_n_closed _ _ _ H)
end.
+ match goal with
H : typed_ctx_item _ _ _ _ _ |- _ => inversion H; subst
end.
eapply bin_log_related_case;
eauto using binary_fundamental;
try match goal with
H : _ |- _ => eapply (typed_n_closed _ _ _ H)
end;
match goal with
H : _ |- _ => eapply (typed_ctx_n_closed _ _ _ _ _ _ _ H)
end.
+ match goal with
H : typed_ctx_item _ _ _ _ _ |- _ => inversion H; subst
end.
eapply bin_log_related_case;
eauto using binary_fundamental;
try match goal with
H : _ |- _ => eapply (typed_n_closed _ _ _ H)
end;
match goal with
H : _ |- _ => eapply (typed_ctx_n_closed _ _ _ _ _ _ _ H)
end.
+ eapply bin_log_related_if;
eauto using typed_ctx_typed, binary_fundamental.
+ eapply bin_log_related_if;
eauto using typed_ctx_typed, binary_fundamental.
+ eapply bin_log_related_if;
eauto using typed_ctx_typed, binary_fundamental.
+ eapply bin_log_related_int_binop;
eauto using typed_ctx_typed, binary_fundamental.
+ eapply bin_log_related_int_binop;
eauto using typed_ctx_typed, binary_fundamental.
+ eapply bin_log_related_fold; eauto.
+ eapply bin_log_related_unfold; eauto.
+ eapply bin_log_related_tlam; eauto with typeclass_instances.
+ eapply bin_log_related_tapp; eauto.
+ eapply bin_log_related_alloc; eauto.
+ eapply bin_log_related_read; eauto.
+ eapply bin_log_related_write; eauto using binary_fundamental.
+ eapply bin_log_related_write; eauto using binary_fundamental.
+ eapply bin_log_related_compare; eauto using binary_fundamental.
+ eapply bin_log_related_compare; eauto using binary_fundamental.
+ eapply bin_log_related_return; eauto using binary_fundamental.
+ eapply bin_log_related_bind; eauto using binary_fundamental.
+ eapply bin_log_related_bind; eauto using binary_fundamental.
+ eapply bin_log_related_runST; eauto using binary_fundamental.
Unshelve. all: trivial.
Qed.
End bin_log_related_under_typed_ctx.
logrelNN fundamentalNN contextual_refinement.
Ltac fold_interp :=
match goal with
| |- context [ interp_exprNN (interp_arrow (interpNN ?A) (interpNN ?A'))
?B (?C, ?D) ] =>
change (interp_exprNN (interp_arrow (interpNN A) (interpNN A')) B (C, D)) with
(interp_exprNN (interpNN (TArrow A A')) B (C, D))
| |- context [ interp_exprNN (interp_prod (interpNN ?A) (interpNN ?A'))
?B (?C, ?D) ] =>
change (interp_exprNN (interp_prod (interpNN A) (interpNN A')) B (C, D)) with
(interp_exprNN (interpNN (TProd A A')) B (C, D))
| |- context [ interp_exprNN (interp_prod (interpNN ?A) (interpNN ?A'))
?B (?C, ?D) ] =>
change (interp_exprNN (interp_prod (interpNN A) (interpNN A')) B (C, D)) with
(interp_exprNN (interpNN (TProd A A')) B (C, D))
| |- context [ interp_exprNN (interp_sum (interpNN ?A) (interpNN ?A'))
?B (?C, ?D) ] =>
change (interp_exprNN (interp_sum (interpNN A) (interpNN A')) B (C, D)) with
(interp_exprNN (interpNN (TSum A A')) B (C, D))
| |- context [ interp_exprNN (interp_rec (interpNN ?A)) ?B (?C, ?D) ] =>
change (interp_exprNN (interp_rec (interpNN A)) B (C, D)) with
(interp_exprNN (interpNN (TRec A)) B (C, D))
| |- context [ interp_exprNN (interp_forall (interpNN ?A))
?B (?C, ?D) ] =>
change (interp_exprNN (interp_forall (interpNN A)) B (C, D)) with
(interp_exprNN (interpNN (TForall A)) B (C, D))
| |- context [ interp_exprNN (interp_ref ?E (interpNN ?A))
?B (?C, ?D) ] =>
change (interp_exprNN (interp_ref E (interpNN A)) B (C, D)) with
(interp_exprNN (interpNN (TSTref E A)) B (C, D))
end.
Section bin_log_related_under_typed_ctx.
Context `{heapG Σ, LogRelG Σ}.
Lemma bin_log_related_under_typed_ctx Γ e e' τ Γ' τ' K :
(∀ f, e.[upn (length Γ) f] = e) →
(∀ f, e'.[upn (length Γ) f] = e') →
typed_ctx K Γ τ Γ' τ' →
Γ ⊨ e ≤logNN≤ e' : τ → Γ' ⊨ fill_ctx K e ≤logNN≤ fill_ctx K e' : τ'.
Proof.
revert Γ τ Γ' τ' e e'.
induction K as [|k K] => Γ τ Γ' τ' e e' H1 H2; simpl.
- inversion_clear 1; trivial.
- inversion_clear 1 as [|? ? ? ? ? ? ? ? Hx1 Hx2]. intros H3.
specialize (IHK _ _ _ _ e e' H1 H2 Hx2 H3).
inversion Hx1; subst; simpl; try fold_interp.
+ eapply bin_log_related_rec; eauto;
match goal with
H : _ |- _ => eapply (typed_ctx_n_closed _ _ _ _ _ _ _ H)
end.
+ eapply bin_log_related_app; eauto using binary_fundamental.
+ eapply bin_log_related_app; eauto using binary_fundamental.
+ eapply bin_log_related_pair; eauto using binary_fundamental.
+ eapply bin_log_related_pair; eauto using binary_fundamental.
+ eapply bin_log_related_fst; eauto.
+ eapply bin_log_related_snd; eauto.
+ eapply bin_log_related_injl; eauto.
+ eapply bin_log_related_injr; eauto.
+ match goal with
H : typed_ctx_item _ _ _ _ _ |- _ => inversion H; subst
end.
eapply bin_log_related_case;
eauto using binary_fundamental;
match goal with
H : _ |- _ => eapply (typed_n_closed _ _ _ H)
end.
+ match goal with
H : typed_ctx_item _ _ _ _ _ |- _ => inversion H; subst
end.
eapply bin_log_related_case;
eauto using binary_fundamental;
try match goal with
H : _ |- _ => eapply (typed_n_closed _ _ _ H)
end;
match goal with
H : _ |- _ => eapply (typed_ctx_n_closed _ _ _ _ _ _ _ H)
end.
+ match goal with
H : typed_ctx_item _ _ _ _ _ |- _ => inversion H; subst
end.
eapply bin_log_related_case;
eauto using binary_fundamental;
try match goal with
H : _ |- _ => eapply (typed_n_closed _ _ _ H)
end;
match goal with
H : _ |- _ => eapply (typed_ctx_n_closed _ _ _ _ _ _ _ H)
end.
+ eapply bin_log_related_if;
eauto using typed_ctx_typed, binary_fundamental.
+ eapply bin_log_related_if;
eauto using typed_ctx_typed, binary_fundamental.
+ eapply bin_log_related_if;
eauto using typed_ctx_typed, binary_fundamental.
+ eapply bin_log_related_int_binop;
eauto using typed_ctx_typed, binary_fundamental.
+ eapply bin_log_related_int_binop;
eauto using typed_ctx_typed, binary_fundamental.
+ eapply bin_log_related_fold; eauto.
+ eapply bin_log_related_unfold; eauto.
+ eapply bin_log_related_tlam; eauto with typeclass_instances.
+ eapply bin_log_related_tapp; eauto.
+ eapply bin_log_related_alloc; eauto.
+ eapply bin_log_related_read; eauto.
+ eapply bin_log_related_write; eauto using binary_fundamental.
+ eapply bin_log_related_write; eauto using binary_fundamental.
+ eapply bin_log_related_compare; eauto using binary_fundamental.
+ eapply bin_log_related_compare; eauto using binary_fundamental.
+ eapply bin_log_related_return; eauto using binary_fundamental.
+ eapply bin_log_related_bind; eauto using binary_fundamental.
+ eapply bin_log_related_bind; eauto using binary_fundamental.
+ eapply bin_log_related_runST; eauto using binary_fundamental.
Unshelve. all: trivial.
Qed.
End bin_log_related_under_typed_ctx.