parent
0538dc6c8a
commit
89fcd6c249
@ -0,0 +1,47 @@
|
||||
From iris.prelude Require Import prelude options.
|
||||
From Autosubst Require Export Autosubst.
|
||||
|
||||
|
||||
(* [idss n sigma] alters the substitution [sigma] by
|
||||
"prepending" binders [0..n), so that [#i] is substituted for binder [i].
|
||||
After that, the binders from [sigma] follow.
|
||||
*)
|
||||
Definition idss `{Ids X} (n : nat) (sigma : var → X) (x : var) :=
|
||||
if decide (x < n) then ids x else sigma (x - n).
|
||||
Lemma idss_0 `{Ids X} (sigma : var → X) :
|
||||
idss 0 sigma = sigma.
|
||||
Proof.
|
||||
f_ext. intros x; unfold idss. simpl.
|
||||
replace (x - 0) with x by lia. done.
|
||||
Qed.
|
||||
Lemma up_idss `{Ids X} `{Rename X} (sigma : var → X) n :
|
||||
(* this precondition holds for the instances we are interested in,
|
||||
but is not a general law assumed by autosubst *)
|
||||
(∀ x : var, rename (+1) (ids x) = ids (S x)) →
|
||||
up (idss n sigma) = idss (S n) (sigma >>> rename (+1) ).
|
||||
Proof.
|
||||
intros Hren_law.
|
||||
f_ext. intros x. destruct x as [ | x]; unfold idss; simpl; first done.
|
||||
unfold up. simpl.
|
||||
destruct (decide (x < n)).
|
||||
- rewrite decide_True; last lia. apply Hren_law.
|
||||
- rewrite decide_False; last lia. done.
|
||||
Qed.
|
||||
|
||||
(* We will also need something like [idss] to shift variable renamings [var → var].
|
||||
Instead of doing a separate definition like
|
||||
[ Definition idsc (n : nat) (sigma : var → var) (x : var) := if decide (x < n) then x else sigma (x - n). ]
|
||||
we declare suitable instances to use [idss] for variable renamings.
|
||||
*)
|
||||
#[global] Instance Ids_var : Ids var. exact id. Defined.
|
||||
#[global] Instance Rename_var : Rename var. exact id. Defined.
|
||||
|
||||
(* a lemma for the nat instance *)
|
||||
Lemma upren_idss sigma n :
|
||||
upren (idss n sigma) = idss (S n) (sigma >>> S).
|
||||
Proof.
|
||||
f_ext. intros x. destruct x as [ | x]; unfold idss; simpl; first done.
|
||||
destruct (decide (x < n)).
|
||||
- rewrite decide_True; [done | lia ].
|
||||
- rewrite decide_False; lia.
|
||||
Qed.
|
@ -0,0 +1,189 @@
|
||||
From semantics.ts.stlc Require Import lang.
|
||||
From stdpp Require Import prelude.
|
||||
From iris Require Import prelude.
|
||||
From semantics.lib Require Export maps .
|
||||
|
||||
(** * Parallel substitution *)
|
||||
|
||||
(** This is the parallel substitution operation. We represent a substitution map as a finite map [xs]. *)
|
||||
Fixpoint subst_map (xs : gmap string expr) (e : expr) : expr :=
|
||||
match e with
|
||||
| LitInt n => LitInt n
|
||||
| Var y => match xs !! y with Some es => es | _ => Var y end
|
||||
| App e1 e2 => App (subst_map xs e1) (subst_map xs e2)
|
||||
| Lam x e => Lam x (subst_map (binder_delete x xs) e)
|
||||
| Plus e1 e2 => Plus (subst_map xs e1) (subst_map xs e2)
|
||||
end.
|
||||
|
||||
Lemma subst_map_empty e :
|
||||
subst_map ∅ e = e.
|
||||
Proof.
|
||||
induction e; simpl; f_equal; eauto.
|
||||
destruct x; simpl; [done | by rewrite !delete_empty..].
|
||||
Qed.
|
||||
|
||||
|
||||
|
||||
Lemma subst_map_closed X e xs :
|
||||
closed X e →
|
||||
(∀ x : string, x ∈ dom xs → x ∉ X) →
|
||||
subst_map xs e = e.
|
||||
Proof.
|
||||
intros Hclosed Hd.
|
||||
induction e in xs, X, Hd, Hclosed |-*; simpl in *;try done.
|
||||
{ (* Var *)
|
||||
apply bool_decide_spec in Hclosed.
|
||||
assert (xs !! x = None) as ->; last done.
|
||||
destruct (xs !! x) as [s | ] eqn:Helem; last done.
|
||||
exfalso; eapply Hd; last apply Hclosed.
|
||||
apply elem_of_dom; eauto.
|
||||
}
|
||||
{ (* lambdas *)
|
||||
erewrite IHe; [done | done |].
|
||||
intros y. destruct x as [ | x]; first apply Hd.
|
||||
simpl.
|
||||
rewrite dom_delete elem_of_difference not_elem_of_singleton.
|
||||
intros [Hnx%Hd Hneq]. rewrite elem_of_cons. intros [? | ?]; done.
|
||||
}
|
||||
(* all other cases *)
|
||||
all: repeat match goal with
|
||||
| H : Is_true (_ && _) |- _ => apply andb_True in H as [ ? ? ]
|
||||
end.
|
||||
all: repeat match goal with
|
||||
| H : ∀ _ _, _ → _ → subst_map _ _ = _ |- _ => erewrite H; clear H
|
||||
end; done.
|
||||
Qed.
|
||||
|
||||
Lemma subst_map_subst map x (e e': expr) :
|
||||
closed [] e' →
|
||||
subst_map map (subst x e' e) = subst_map (<[x:=e']>map) e.
|
||||
Proof.
|
||||
intros He'; induction e as [y|y e IH | | |]in map|-*; simpl; try (f_equal; eauto).
|
||||
- case_decide.
|
||||
+ simplify_eq/=. rewrite lookup_insert.
|
||||
rewrite (subst_map_closed []); [done | apply He' | ].
|
||||
intros ? ?. apply not_elem_of_nil.
|
||||
+ rewrite lookup_insert_ne; done.
|
||||
- destruct y; simpl; first done.
|
||||
+ case_decide.
|
||||
* simplify_eq/=. by rewrite delete_insert_delete.
|
||||
* rewrite delete_insert_ne; last by congruence. done.
|
||||
Qed.
|
||||
|
||||
(** We lift the notion of closedness [closed] to substitution maps. *)
|
||||
Definition subst_closed (X : list string) (map : gmap string expr) :=
|
||||
∀ x e, map !! x = Some e → closed X e.
|
||||
Lemma subst_closed_subseteq X map1 map2 :
|
||||
map1 ⊆ map2 → subst_closed X map2 → subst_closed X map1.
|
||||
Proof.
|
||||
intros Hsub Hclosed2 x e Hl. eapply Hclosed2, map_subseteq_spec; done.
|
||||
Qed.
|
||||
|
||||
Lemma subst_closed_weaken X Y map1 map2 :
|
||||
Y ⊆ X → map1 ⊆ map2 → subst_closed Y map2 → subst_closed X map1.
|
||||
Proof.
|
||||
intros Hsub1 Hsub2 Hclosed2 x e Hl.
|
||||
eapply closed_weaken. 1:eapply Hclosed2, map_subseteq_spec; done. done.
|
||||
Qed.
|
||||
|
||||
(** Lemma about the interaction with "normal" substitution. *)
|
||||
Lemma subst_subst_map x es map e :
|
||||
subst_closed [] map →
|
||||
subst x es (subst_map (delete x map) e) =
|
||||
subst_map (<[x:=es]> map) e.
|
||||
Proof.
|
||||
revert map es x; induction e; intros map v0 xx Hclosed; simpl;
|
||||
try (f_equal; eauto).
|
||||
- destruct (decide (xx=x)) as [->|Hne].
|
||||
+ rewrite lookup_delete // lookup_insert //. simpl.
|
||||
rewrite decide_True //.
|
||||
+ rewrite lookup_delete_ne // lookup_insert_ne //.
|
||||
destruct (_ !! x) as [rr|] eqn:Helem.
|
||||
* apply Hclosed in Helem.
|
||||
by apply subst_closed_nil.
|
||||
* simpl. rewrite decide_False //.
|
||||
- destruct x; simpl; first by auto.
|
||||
case_decide.
|
||||
+ simplify_eq. rewrite delete_idemp delete_insert_delete. done.
|
||||
+ rewrite delete_insert_ne //; last congruence. rewrite delete_commute. apply IHe.
|
||||
eapply subst_closed_subseteq; last done.
|
||||
apply map_delete_subseteq.
|
||||
Qed.
|
||||
|
||||
Lemma subst'_subst_map b (es : expr) map e :
|
||||
subst_closed [] map →
|
||||
subst' b es (subst_map (binder_delete b map) e) =
|
||||
subst_map (binder_insert b es map) e.
|
||||
Proof.
|
||||
destruct b; first done.
|
||||
apply subst_subst_map.
|
||||
Qed.
|
||||
|
||||
Lemma closed_subst_weaken e map X Y :
|
||||
subst_closed [] map →
|
||||
(∀ x, x ∈ X → x ∉ dom map → x ∈ Y) →
|
||||
closed X e →
|
||||
closed Y (subst_map map e).
|
||||
Proof.
|
||||
induction e in X, Y, map |-*; simpl; intros Hmclosed Hsub Hcl.
|
||||
{ (* vars *)
|
||||
destruct (map !! x) as [es | ] eqn:Heq.
|
||||
+ apply closed_weaken_nil. by eapply Hmclosed.
|
||||
+ apply bool_decide_pack. apply Hsub; first by eapply bool_decide_unpack.
|
||||
by apply not_elem_of_dom.
|
||||
}
|
||||
{ (* lambdas *)
|
||||
eapply IHe; last done.
|
||||
+ eapply subst_closed_subseteq; last done.
|
||||
destruct x; first done. apply map_delete_subseteq.
|
||||
+ intros y. destruct x as [ | x]; first by apply Hsub.
|
||||
rewrite !elem_of_cons. intros [-> | Hy] Hn; first by left.
|
||||
destruct (decide (y = x)) as [ -> | Hneq]; first by left.
|
||||
right. eapply Hsub; first done. set_solver.
|
||||
}
|
||||
(* all other cases *)
|
||||
all: repeat match goal with
|
||||
| H : Is_true (_ && _) |- _ => apply andb_True in H as [ ? ? ]
|
||||
end.
|
||||
all: repeat match goal with
|
||||
| |- Is_true (_ && _) => apply andb_True; split
|
||||
end.
|
||||
all: naive_solver.
|
||||
Qed.
|
||||
|
||||
|
||||
Lemma subst_map_closed' X Y Θ e:
|
||||
closed Y e →
|
||||
(∀ x, x ∈ Y → if Θ !! x is (Some e') then closed X e' else x ∈ X) →
|
||||
closed X (subst_map Θ e).
|
||||
Proof.
|
||||
induction e in X, Θ, Y |-*; simpl.
|
||||
- intros Hel%bool_decide_unpack Hcl.
|
||||
eapply Hcl in Hel.
|
||||
destruct (Θ !! x); first done.
|
||||
simpl. by eapply bool_decide_pack.
|
||||
- intros Hcl Hcl'. destruct x as [|x]; simpl; first naive_solver.
|
||||
eapply IHe; first done.
|
||||
intros y [|]%elem_of_cons.
|
||||
+ subst. rewrite lookup_delete. set_solver.
|
||||
+ destruct (decide (x = y)); first by subst; rewrite lookup_delete; set_solver.
|
||||
rewrite lookup_delete_ne //=. eapply Hcl' in H.
|
||||
destruct lookup; last set_solver.
|
||||
eapply closed_weaken; eauto with set_solver.
|
||||
- rewrite !andb_True. intros [H1 H2] Hcl. split; eauto.
|
||||
- auto.
|
||||
- rewrite !andb_True. intros [H1 H2] Hcl. split; eauto.
|
||||
Qed.
|
||||
|
||||
Lemma subst_map_closed'_2 X Θ e:
|
||||
closed (X ++ (elements (dom Θ))) e ->
|
||||
subst_closed X Θ ->
|
||||
closed X (subst_map Θ e).
|
||||
Proof.
|
||||
intros Hcl Hsubst.
|
||||
eapply subst_map_closed'; first eassumption.
|
||||
intros x Hx.
|
||||
destruct (Θ !! x) as [e'|] eqn:Heq.
|
||||
- eauto.
|
||||
- by eapply elem_of_app in Hx as [H|H%elem_of_elements%not_elem_of_dom].
|
||||
Qed.
|
@ -0,0 +1,193 @@
|
||||
From stdpp Require Import prelude.
|
||||
From iris Require Import prelude.
|
||||
From semantics.ts.stlc_extended Require Import lang.
|
||||
From semantics.lib Require Import maps.
|
||||
|
||||
Fixpoint subst_map (xs : gmap string expr) (e : expr) : expr :=
|
||||
match e with
|
||||
| LitInt n => LitInt n
|
||||
| Var y => match xs !! y with Some es => es | _ => Var y end
|
||||
| App e1 e2 => App (subst_map xs e1) (subst_map xs e2)
|
||||
| Lam x e => Lam x (subst_map (binder_delete x xs) e)
|
||||
| Plus e1 e2 => Plus (subst_map xs e1) (subst_map xs e2)
|
||||
| Pair e1 e2 => Pair (subst_map xs e1) (subst_map xs e2)
|
||||
| Fst e => Fst (subst_map xs e)
|
||||
| Snd e => Snd (subst_map xs e)
|
||||
| InjL e => InjL (subst_map xs e)
|
||||
| InjR e => InjR (subst_map xs e)
|
||||
| Case e e1 e2 => Case (subst_map xs e) (subst_map xs e1) (subst_map xs e2)
|
||||
end.
|
||||
|
||||
|
||||
Lemma subst_map_empty e :
|
||||
subst_map ∅ e = e.
|
||||
Proof.
|
||||
induction e; simpl; f_equal; eauto.
|
||||
destruct x; simpl; [done | by rewrite !delete_empty..].
|
||||
Qed.
|
||||
|
||||
|
||||
|
||||
Lemma subst_map_closed X e xs :
|
||||
closed X e →
|
||||
(∀ x : string, x ∈ dom xs → x ∉ X) →
|
||||
subst_map xs e = e.
|
||||
Proof.
|
||||
intros Hclosed Hd.
|
||||
induction e in xs, X, Hd, Hclosed |-*; simpl in *;try done.
|
||||
{ (* Var *)
|
||||
apply bool_decide_spec in Hclosed.
|
||||
assert (xs !! x = None) as ->; last done.
|
||||
destruct (xs !! x) as [s | ] eqn:Helem; last done.
|
||||
exfalso; eapply Hd; last apply Hclosed.
|
||||
apply elem_of_dom; eauto.
|
||||
}
|
||||
{ (* lambdas *)
|
||||
erewrite IHe; [done | done |].
|
||||
intros y. destruct x as [ | x]; first apply Hd.
|
||||
simpl.
|
||||
rewrite dom_delete elem_of_difference not_elem_of_singleton.
|
||||
intros [Hnx%Hd Hneq]. rewrite elem_of_cons. intros [? | ?]; done.
|
||||
}
|
||||
(* all other cases *)
|
||||
all: unfold closed in *; simpl in *.
|
||||
all: repeat match goal with
|
||||
| H : Is_true (_ && _) |- _ => apply andb_True in H as [ ? ? ]
|
||||
end.
|
||||
all: repeat match goal with
|
||||
| H : ∀ _ _, _ → _ → subst_map _ _ = _ |- _ => erewrite H; clear H
|
||||
end; try done.
|
||||
Qed.
|
||||
|
||||
Lemma subst_map_subst map x (e e': expr) :
|
||||
closed [] e' →
|
||||
subst_map map (subst x e' e) = subst_map (<[x:=e']>map) e.
|
||||
Proof.
|
||||
intros He'; induction e as [y|y e IH | | | | | | | | | ]in map|-*; simpl; try (f_equal; eauto).
|
||||
- case_decide.
|
||||
+ simplify_eq/=. rewrite lookup_insert.
|
||||
rewrite (subst_map_closed []); [done | apply He' | ].
|
||||
intros ? ?. apply not_elem_of_nil.
|
||||
+ rewrite lookup_insert_ne; done.
|
||||
- destruct y; simpl; first done.
|
||||
+ case_decide.
|
||||
* simplify_eq/=. by rewrite delete_insert_delete.
|
||||
* rewrite delete_insert_ne; last by congruence. done.
|
||||
Qed.
|
||||
|
||||
(** We lift the notion of closedness [closed] to substitution maps. *)
|
||||
Definition subst_closed (X : list string) (map : gmap string expr) :=
|
||||
∀ x e, map !! x = Some e → closed X e.
|
||||
Lemma subst_closed_subseteq X map1 map2 :
|
||||
map1 ⊆ map2 → subst_closed X map2 → subst_closed X map1.
|
||||
Proof.
|
||||
intros Hsub Hclosed2 x e Hl. eapply Hclosed2, map_subseteq_spec; done.
|
||||
Qed.
|
||||
|
||||
Lemma subst_closed_weaken X Y map1 map2 :
|
||||
Y ⊆ X → map1 ⊆ map2 → subst_closed Y map2 → subst_closed X map1.
|
||||
Proof.
|
||||
intros Hsub1 Hsub2 Hclosed2 x e Hl.
|
||||
eapply is_closed_weaken. 1:eapply Hclosed2, map_subseteq_spec; done. done.
|
||||
Qed.
|
||||
|
||||
(** Lemma about the interaction with "normal" substitution. *)
|
||||
Lemma subst_subst_map x es map e :
|
||||
subst_closed [] map →
|
||||
subst x es (subst_map (delete x map) e) =
|
||||
subst_map (<[x:=es]> map) e.
|
||||
Proof.
|
||||
revert map es x; induction e; intros map v0 xx Hclosed; simpl;
|
||||
try (f_equal; eauto).
|
||||
- destruct (decide (xx=x)) as [->|Hne].
|
||||
+ rewrite lookup_delete // lookup_insert //. simpl.
|
||||
rewrite decide_True //.
|
||||
+ rewrite lookup_delete_ne // lookup_insert_ne //.
|
||||
destruct (_ !! x) as [rr|] eqn:Helem.
|
||||
* apply Hclosed in Helem.
|
||||
by apply subst_is_closed_nil.
|
||||
* simpl. rewrite decide_False //.
|
||||
- destruct x; simpl; first by auto.
|
||||
case_decide.
|
||||
+ simplify_eq. rewrite delete_idemp delete_insert_delete. done.
|
||||
+ rewrite delete_insert_ne //; last congruence. rewrite delete_commute. apply IHe.
|
||||
eapply subst_closed_subseteq; last done.
|
||||
apply map_delete_subseteq.
|
||||
Qed.
|
||||
|
||||
Lemma subst'_subst_map b (es : expr) map e :
|
||||
subst_closed [] map →
|
||||
subst' b es (subst_map (binder_delete b map) e) =
|
||||
subst_map (binder_insert b es map) e.
|
||||
Proof.
|
||||
destruct b; first done.
|
||||
apply subst_subst_map.
|
||||
Qed.
|
||||
|
||||
Lemma closed_subst_weaken e map X Y :
|
||||
subst_closed [] map →
|
||||
(∀ x, x ∈ X → x ∉ dom map → x ∈ Y) →
|
||||
closed X e →
|
||||
closed Y (subst_map map e).
|
||||
Proof.
|
||||
induction e in X, Y, map |-*; simpl; intros Hmclosed Hsub Hcl.
|
||||
{ (* vars *)
|
||||
destruct (map !! x) as [es | ] eqn:Heq.
|
||||
+ apply is_closed_weaken_nil. by eapply Hmclosed.
|
||||
+ apply bool_decide_pack. apply Hsub; first by eapply bool_decide_unpack.
|
||||
by apply not_elem_of_dom.
|
||||
}
|
||||
{ (* lambdas *)
|
||||
eapply IHe; last done.
|
||||
+ eapply subst_closed_subseteq; last done.
|
||||
destruct x; first done. apply map_delete_subseteq.
|
||||
+ intros y. destruct x as [ | x]; first by apply Hsub.
|
||||
rewrite !elem_of_cons. intros [-> | Hy] Hn; first by left.
|
||||
destruct (decide (y = x)) as [ -> | Hneq]; first by left.
|
||||
right. eapply Hsub; first done. set_solver.
|
||||
}
|
||||
(* all other cases *)
|
||||
all: unfold closed in *; simpl in *.
|
||||
all: repeat match goal with
|
||||
| H : Is_true (_ && _) |- _ => apply andb_True in H as [ ? ? ]
|
||||
end.
|
||||
all: repeat match goal with
|
||||
| |- Is_true (_ && _) => apply andb_True; split
|
||||
end.
|
||||
all: try naive_solver.
|
||||
Qed.
|
||||
|
||||
|
||||
Lemma subst_map_closed' X Y Θ e:
|
||||
closed Y e →
|
||||
(∀ x, x ∈ Y → if Θ !! x is (Some e') then closed X e' else x ∈ X) →
|
||||
closed X (subst_map Θ e).
|
||||
Proof.
|
||||
induction e in X, Θ, Y |-*; simpl.
|
||||
{ intros Hel%bool_decide_unpack Hcl.
|
||||
eapply Hcl in Hel.
|
||||
destruct (Θ !! x); first done.
|
||||
simpl. by eapply bool_decide_pack. }
|
||||
{ intros Hcl Hcl'. destruct x as [|x]; simpl; first naive_solver.
|
||||
eapply IHe; first done.
|
||||
intros y [|]%elem_of_cons.
|
||||
+ subst. rewrite lookup_delete. set_solver.
|
||||
+ destruct (decide (x = y)); first by subst; rewrite lookup_delete; set_solver.
|
||||
rewrite lookup_delete_ne //=. eapply Hcl' in H.
|
||||
destruct lookup; last set_solver.
|
||||
eapply is_closed_weaken; eauto with set_solver. }
|
||||
all: unfold closed; simpl; naive_solver.
|
||||
Qed.
|
||||
|
||||
Lemma subst_map_closed'_2 X Θ e:
|
||||
closed (X ++ (elements (dom Θ))) e ->
|
||||
subst_closed X Θ ->
|
||||
closed X (subst_map Θ e).
|
||||
Proof.
|
||||
intros Hcl Hsubst.
|
||||
eapply subst_map_closed'; first eassumption.
|
||||
intros x Hx.
|
||||
destruct (Θ !! x) as [e'|] eqn:Heq.
|
||||
- eauto.
|
||||
- by eapply elem_of_app in Hx as [H|H%elem_of_elements%not_elem_of_dom].
|
||||
Qed.
|
@ -0,0 +1,100 @@
|
||||
From stdpp Require Import gmap base relations.
|
||||
From iris Require Import prelude.
|
||||
From semantics.ts.systemf Require Import lang notation types tactics.
|
||||
|
||||
(** Exercise 3 (LN Exercise 22): Universal Fun *)
|
||||
|
||||
Definition fun_comp : val :=
|
||||
#0 (* TODO *).
|
||||
Definition fun_comp_type : type :=
|
||||
#0 (* TODO *).
|
||||
Lemma fun_comp_typed :
|
||||
TY 0; ∅ ⊢ fun_comp : fun_comp_type.
|
||||
Proof.
|
||||
(* should be solved by solve_typing. *)
|
||||
(* TODO: exercise *)
|
||||
Admitted.
|
||||
|
||||
|
||||
Definition swap_args : val :=
|
||||
#0 (* TODO *).
|
||||
Definition swap_args_type : type :=
|
||||
#0 (* TODO *).
|
||||
Lemma swap_args_typed :
|
||||
TY 0; ∅ ⊢ swap_args : swap_args_type.
|
||||
Proof.
|
||||
(* should be solved by solve_typing. *)
|
||||
(* TODO: exercise *)
|
||||
Admitted.
|
||||
|
||||
|
||||
Definition lift_prod : val :=
|
||||
#0 (* TODO *).
|
||||
Definition lift_prod_type : type :=
|
||||
#0 (* TODO *).
|
||||
Lemma lift_prod_typed :
|
||||
TY 0; ∅ ⊢ lift_prod : lift_prod_type.
|
||||
Proof.
|
||||
(* should be solved by solve_typing. *)
|
||||
(* TODO: exercise *)
|
||||
Admitted.
|
||||
|
||||
|
||||
Definition lift_sum : val :=
|
||||
#0 (* TODO *).
|
||||
Definition lift_sum_type : type :=
|
||||
#0 (* TODO *).
|
||||
Lemma lift_sum_typed :
|
||||
TY 0; ∅ ⊢ lift_sum : lift_sum_type.
|
||||
Proof.
|
||||
(* should be solved by solve_typing. *)
|
||||
(* TODO: exercise *)
|
||||
Admitted.
|
||||
|
||||
|
||||
(** Exercise 5 (LN Exercise 18): Named to De Bruijn *)
|
||||
Inductive ptype : Type :=
|
||||
| PTVar : string → ptype
|
||||
| PInt
|
||||
| PBool
|
||||
| PTForall : string → ptype → ptype
|
||||
| PTExists : string → ptype → ptype
|
||||
| PFun (A B : ptype).
|
||||
|
||||
Declare Scope PType_scope.
|
||||
Delimit Scope PType_scope with pty.
|
||||
Bind Scope PType_scope with ptype.
|
||||
Coercion PTVar: string >-> ptype.
|
||||
Infix "→" := PFun : PType_scope.
|
||||
Notation "∀: x , τ" :=
|
||||
(PTForall x τ%pty)
|
||||
(at level 100, τ at level 200) : PType_scope.
|
||||
Notation "∃: x , τ" :=
|
||||
(PTExists x τ%pty)
|
||||
(at level 100, τ at level 200) : PType_scope.
|
||||
|
||||
Fixpoint debruijn (m: gmap string nat) (A: ptype) : option type :=
|
||||
None (* FIXME *).
|
||||
|
||||
(* Example *)
|
||||
Goal debruijn ∅ (∀: "x", ∀: "y", "x" → "y")%pty = Some (∀: ∀: #1 → #0)%ty.
|
||||
Proof.
|
||||
(* Should be solved by reflexivity. *)
|
||||
(* TODO: exercise *)
|
||||
Admitted.
|
||||
|
||||
|
||||
Goal debruijn ∅ (∀: "x", "x" → ∀: "y", "y")%pty = Some (∀: #0 → ∀: #0)%ty.
|
||||
Proof.
|
||||
(* Should be solved by reflexivity. *)
|
||||
(* TODO: exercise *)
|
||||
Admitted.
|
||||
|
||||
|
||||
Goal debruijn ∅ (∀: "x", "x" → ∀: "y", "x")%pty = Some (∀: #0 → ∀: #1)%ty.
|
||||
Proof.
|
||||
(* Should be solved by reflexivity. *)
|
||||
(* TODO: exercise *)
|
||||
Admitted.
|
||||
|
||||
|
@ -0,0 +1,516 @@
|
||||
From stdpp Require Export binders strings.
|
||||
From iris.prelude Require Import options.
|
||||
From semantics.lib Require Export maps.
|
||||
|
||||
Declare Scope expr_scope.
|
||||
Declare Scope val_scope.
|
||||
Delimit Scope expr_scope with E.
|
||||
Delimit Scope val_scope with V.
|
||||
|
||||
(** Expressions and vals. *)
|
||||
Inductive base_lit : Set :=
|
||||
| LitInt (n : Z) | LitBool (b : bool) | LitUnit.
|
||||
Inductive un_op : Set :=
|
||||
| NegOp | MinusUnOp.
|
||||
Inductive bin_op : Set :=
|
||||
| PlusOp | MinusOp | MultOp (* Arithmetic *)
|
||||
| LtOp | LeOp | EqOp. (* Comparison *)
|
||||
|
||||
|
||||
Inductive expr :=
|
||||
| Lit (l : base_lit)
|
||||
(* Base lambda calculus *)
|
||||
| Var (x : string)
|
||||
| Lam (x : binder) (e : expr)
|
||||
| App (e1 e2 : expr)
|
||||
(* Base types and their operations *)
|
||||
| UnOp (op : un_op) (e : expr)
|
||||
| BinOp (op : bin_op) (e1 e2 : expr)
|
||||
| If (e0 e1 e2 : expr)
|
||||
(* Polymorphism *)
|
||||
| TApp (e : expr)
|
||||
| TLam (e : expr)
|
||||
| Pack (e : expr)
|
||||
| Unpack (x : binder) (e1 e2 : expr)
|
||||
(* Products *)
|
||||
| Pair (e1 e2 : expr)
|
||||
| Fst (e : expr)
|
||||
| Snd (e : expr)
|
||||
(* Sums *)
|
||||
| InjL (e : expr)
|
||||
| InjR (e : expr)
|
||||
| Case (e0 : expr) (e1 : expr) (e2 : expr).
|
||||
|
||||
Bind Scope expr_scope with expr.
|
||||
|
||||
Inductive val :=
|
||||
| LitV (l : base_lit)
|
||||
| LamV (x : binder) (e : expr)
|
||||
| TLamV (e : expr)
|
||||
| PackV (v : val)
|
||||
| PairV (v1 v2 : val)
|
||||
| InjLV (v : val)
|
||||
| InjRV (v : val)
|
||||
.
|
||||
|
||||
Bind Scope val_scope with val.
|
||||
|
||||
Fixpoint of_val (v : val) : expr :=
|
||||
match v with
|
||||
| LitV l => Lit l
|
||||
| LamV x e => Lam x e
|
||||
| TLamV e => TLam e
|
||||
| PackV v => Pack (of_val v)
|
||||
| PairV v1 v2 => Pair (of_val v1) (of_val v2)
|
||||
| InjLV v => InjL (of_val v)
|
||||
| InjRV v => InjR (of_val v)
|
||||
end.
|
||||
|
||||
Fixpoint to_val (e : expr) : option val :=
|
||||
match e with
|
||||
| Lit l => Some $ LitV l
|
||||
| Lam x e => Some (LamV x e)
|
||||
| TLam e => Some (TLamV e)
|
||||
| Pack e =>
|
||||
to_val e ≫= (λ v, Some $ PackV v)
|
||||
| Pair e1 e2 =>
|
||||
to_val e1 ≫= (λ v1, to_val e2 ≫= (λ v2, Some $ PairV v1 v2))
|
||||
| InjL e => to_val e ≫= (λ v, Some $ InjLV v)
|
||||
| InjR e => to_val e ≫= (λ v, Some $ InjRV v)
|
||||
| _ => None
|
||||
end.
|
||||
|
||||
(** Equality and other typeclass stuff *)
|
||||
Lemma to_of_val v : to_val (of_val v) = Some v.
|
||||
Proof.
|
||||
by induction v; simplify_option_eq; repeat f_equal; try apply (proof_irrel _).
|
||||
Qed.
|
||||
|
||||
Lemma of_to_val e v : to_val e = Some v → of_val v = e.
|
||||
Proof.
|
||||
revert v; induction e; intros v ?; simplify_option_eq; auto with f_equal.
|
||||
Qed.
|
||||
|
||||
#[export] Instance of_val_inj : Inj (=) (=) of_val.
|
||||
Proof. by intros ?? Hv; apply (inj Some); rewrite <-!to_of_val, Hv. Qed.
|
||||
|
||||
(** structural computational version *)
|
||||
Fixpoint is_val (e : expr) : Prop :=
|
||||
match e with
|
||||
| Lit l => True
|
||||
| Lam x e => True
|
||||
| TLam e => True
|
||||
| Pack e => is_val e
|
||||
| Pair e1 e2 => is_val e1 ∧ is_val e2
|
||||
| InjL e => is_val e
|
||||
| InjR e => is_val e
|
||||
| _ => False
|
||||
end.
|
||||
Lemma is_val_spec e : is_val e ↔ ∃ v, to_val e = Some v.
|
||||
Proof.
|
||||
induction e as [ | | ? e IH | e1 IH1 e2 IH2 | e IH | ? e1 IH1 e2 IH2 | e1 IH1 e2 IH2 e3 IH3 | e IH | e IH | e IH | ? e1 IH1 e2 IH2 | e1 IH1 e2 IH2 | e IH | e IH | e IH | e IH | e1 IH1 e2 IH2 e3 IH3 ];
|
||||
simpl; (split; [ | intros (v & Heq)]); simplify_option_eq; try done; eauto.
|
||||
- rewrite IH. intros (v & ->). eauto.
|
||||
- apply IH. eauto.
|
||||
- rewrite IH1, IH2. intros [(v1 & ->) (v2 & ->)]. eauto.
|
||||
- rewrite IH1, IH2. eauto.
|
||||
- rewrite IH. intros (v & ->). eauto.
|
||||
- apply IH. eauto.
|
||||
- rewrite IH. intros (v & ->); eauto.
|
||||
- apply IH. eauto.
|
||||
Qed.
|
||||
|
||||
Ltac simplify_val :=
|
||||
repeat match goal with
|
||||
| H: to_val (of_val ?v) = ?o |- _ => rewrite to_of_val in H
|
||||
| H: is_val ?e |- _ => destruct (proj1 (is_val_spec e) H) as (? & ?); clear H
|
||||
end.
|
||||
|
||||
(* Misc *)
|
||||
Lemma is_val_of_val v : is_val (of_val v).
|
||||
Proof. apply is_val_spec. rewrite to_of_val. eauto. Qed.
|
||||
|
||||
(* literals and our operators have decidable equality. *)
|
||||
Global Instance base_lit_eq_dec : EqDecision base_lit.
|
||||
Proof. solve_decision. Defined.
|
||||
Global Instance un_op_eq_dec : EqDecision un_op.
|
||||
Proof. solve_decision. Defined.
|
||||
Global Instance bin_op_eq_dec : EqDecision bin_op.
|
||||
Proof. solve_decision. Defined.
|
||||
|
||||
|
||||
(** Substitution *)
|
||||
Fixpoint subst (x : string) (es : expr) (e : expr) : expr :=
|
||||
match e with
|
||||
| Lit _ => e
|
||||
| Var y => if decide (x = y) then es else Var y
|
||||
| Lam y e =>
|
||||
Lam y $ if decide (BNamed x = y) then e else subst x es e
|
||||
| App e1 e2 => App (subst x es e1) (subst x es e2)
|
||||
| TApp e => TApp (subst x es e)
|
||||
| TLam e => TLam (subst x es e)
|
||||
| Pack e => Pack (subst x es e)
|
||||
| Unpack y e1 e2 => Unpack y (subst x es e1) (if decide (BNamed x = y) then e2 else subst x es e2)
|
||||
| UnOp op e => UnOp op (subst x es e)
|
||||
| BinOp op e1 e2 => BinOp op (subst x es e1) (subst x es e2)
|
||||
| If e0 e1 e2 => If (subst x es e0) (subst x es e1) (subst x es e2)
|
||||
| Pair e1 e2 => Pair (subst x es e1) (subst x es e2)
|
||||
| Fst e => Fst (subst x es e)
|
||||
| Snd e => Snd (subst x es e)
|
||||
| InjL e => InjL (subst x es e)
|
||||
| InjR e => InjR (subst x es e)
|
||||
| Case e0 e1 e2 => Case (subst x es e0) (subst x es e1) (subst x es e2)
|
||||
end.
|
||||
|
||||
Definition subst' (mx : binder) (es : expr) : expr → expr :=
|
||||
match mx with BNamed x => subst x es | BAnon => id end.
|
||||
|
||||
(** The stepping relation *)
|
||||
Definition un_op_eval (op : un_op) (v : val) : option val :=
|
||||
match op, v with
|
||||
| NegOp, LitV (LitBool b) => Some $ LitV $ LitBool (negb b)
|
||||
| MinusUnOp, LitV (LitInt n) => Some $ LitV $ LitInt (- n)
|
||||
| _, _ => None
|
||||
end.
|
||||
|
||||
Definition bin_op_eval_int (op : bin_op) (n1 n2 : Z) : option base_lit :=
|
||||
match op with
|
||||
| PlusOp => Some $ LitInt (n1 + n2)
|
||||
| MinusOp => Some $ LitInt (n1 - n2)
|
||||
| MultOp => Some $ LitInt (n1 * n2)
|
||||
| LtOp => Some $ LitBool (bool_decide (n1 < n2))
|
||||
| LeOp => Some $ LitBool (bool_decide (n1 ≤ n2))
|
||||
| EqOp => Some $ LitBool (bool_decide (n1 = n2))
|
||||
end%Z.
|
||||
|
||||
Definition bin_op_eval (op : bin_op) (v1 v2 : val) : option val :=
|
||||
match v1, v2 with
|
||||
| LitV (LitInt n1), LitV (LitInt n2) => LitV <$> bin_op_eval_int op n1 n2
|
||||
| _, _ => None
|
||||
end.
|
||||
|
||||
Inductive base_step : expr → expr → Prop :=
|
||||
| BetaS x e1 e2 e' :
|
||||
is_val e2 →
|
||||
e' = subst' x e2 e1 →
|
||||
base_step (App (Lam x e1) e2) e'
|
||||
| TBetaS e1 :
|
||||
base_step (TApp (TLam e1)) e1
|
||||
| UnpackS e1 e2 e' x :
|
||||
is_val e1 →
|
||||
e' = subst' x e1 e2 →
|
||||
base_step (Unpack x (Pack e1) e2) e'
|
||||
| UnOpS op e v v' :
|
||||
to_val e = Some v →
|
||||
un_op_eval op v = Some v' →
|
||||
base_step (UnOp op e) (of_val v')
|
||||
| BinOpS op e1 v1 e2 v2 v' :
|
||||
to_val e1 = Some v1 →
|
||||
to_val e2 = Some v2 →
|
||||
bin_op_eval op v1 v2 = Some v' →
|
||||
base_step (BinOp op e1 e2) (of_val v')
|
||||
| IfTrueS e1 e2 :
|
||||
base_step (If (Lit $ LitBool true) e1 e2) e1
|
||||
| IfFalseS e1 e2 :
|
||||
base_step (If (Lit $ LitBool false) e1 e2) e2
|
||||
| FstS e1 e2 :
|
||||
is_val e1 →
|
||||
is_val e2 →
|
||||
base_step (Fst (Pair e1 e2)) e1
|
||||
| SndS e1 e2 :
|
||||
is_val e1 →
|
||||
is_val e2 →
|
||||
base_step (Snd (Pair e1 e2)) e2
|
||||
| CaseLS e e1 e2 :
|
||||
is_val e →
|
||||
base_step (Case (InjL e) e1 e2) (App e1 e)
|
||||
| CaseRS e e1 e2 :
|
||||
is_val e →
|
||||
base_step (Case (InjR e) e1 e2) (App e2 e)
|
||||
.
|
||||
|
||||
(** We define evaluation contexts *)
|
||||
Inductive ectx :=
|
||||
| HoleCtx
|
||||
| AppLCtx (K: ectx) (v2 : val)
|
||||
| AppRCtx (e1 : expr) (K: ectx)
|
||||
| TAppCtx (K: ectx)
|
||||
| PackCtx (K: ectx)
|
||||
| UnpackCtx (x : binder)(K: ectx) (e2 : expr)
|
||||
| UnOpCtx (op : un_op) (K: ectx)
|
||||
| BinOpLCtx (op : bin_op) (K: ectx) (v2 : val)
|
||||
| BinOpRCtx (op : bin_op) (e1 : expr) (K: ectx)
|
||||
| IfCtx (K: ectx) (e1 e2 : expr)
|
||||
| PairLCtx (K: ectx) (v2 : val)
|
||||
| PairRCtx (e1 : expr) (K: ectx)
|
||||
| FstCtx (K: ectx)
|
||||
| SndCtx (K: ectx)
|
||||
| InjLCtx (K: ectx)
|
||||
| InjRCtx (K: ectx)
|
||||
| CaseCtx (K: ectx) (e1 : expr) (e2 : expr) .
|
||||
|
||||
Fixpoint fill (K : ectx) (e : expr) : expr :=
|
||||
match K with
|
||||
| HoleCtx => e
|
||||
| AppLCtx K v2 => App (fill K e) (of_val v2)
|
||||
| AppRCtx e1 K => App e1 (fill K e)
|
||||
| TAppCtx K => TApp (fill K e)
|
||||
| PackCtx K => Pack (fill K e)
|
||||
| UnpackCtx x K e2 => Unpack x (fill K e) e2
|
||||
| UnOpCtx op K => UnOp op (fill K e)
|
||||
| BinOpLCtx op K v2 => BinOp op (fill K e) (of_val v2)
|
||||
| BinOpRCtx op e1 K => BinOp op e1 (fill K e)
|
||||
| IfCtx K e1 e2 => If (fill K e) e1 e2
|
||||
| PairLCtx K v2 => Pair (fill K e) (of_val v2)
|
||||
| PairRCtx e1 K => Pair e1 (fill K e)
|
||||
| FstCtx K => Fst (fill K e)
|
||||
| SndCtx K => Snd (fill K e)
|
||||
| InjLCtx K => InjL (fill K e)
|
||||
| InjRCtx K => InjR (fill K e)
|
||||
| CaseCtx K e1 e2 => Case (fill K e) e1 e2
|
||||
end.
|
||||
|
||||
Fixpoint comp_ectx (K: ectx) (K' : ectx) : ectx :=
|
||||
match K with
|
||||
| HoleCtx => K'
|
||||
| AppLCtx K v2 => AppLCtx (comp_ectx K K') v2
|
||||
| AppRCtx e1 K => AppRCtx e1 (comp_ectx K K')
|
||||
| TAppCtx K => TAppCtx (comp_ectx K K')
|
||||
| PackCtx K => PackCtx (comp_ectx K K')
|
||||
| UnpackCtx x K e2 => UnpackCtx x (comp_ectx K K') e2
|
||||
| UnOpCtx op K => UnOpCtx op (comp_ectx K K')
|
||||
| BinOpLCtx op K v2 => BinOpLCtx op (comp_ectx K K') v2
|
||||
| BinOpRCtx op e1 K => BinOpRCtx op e1 (comp_ectx K K')
|
||||
| IfCtx K e1 e2 => IfCtx (comp_ectx K K') e1 e2
|
||||
| PairLCtx K v2 => PairLCtx (comp_ectx K K') v2
|
||||
| PairRCtx e1 K => PairRCtx e1 (comp_ectx K K')
|
||||
| FstCtx K => FstCtx (comp_ectx K K')
|
||||
| SndCtx K => SndCtx (comp_ectx K K')
|
||||
| InjLCtx K => InjLCtx (comp_ectx K K')
|
||||
| InjRCtx K => InjRCtx (comp_ectx K K')
|
||||
| CaseCtx K e1 e2 => CaseCtx (comp_ectx K K') e1 e2
|
||||
end.
|
||||
|
||||
(** Contextual steps *)
|
||||
Inductive contextual_step (e1 : expr) (e2 : expr) : Prop :=
|
||||
Ectx_step K e1' e2' :
|
||||
e1 = fill K e1' → e2 = fill K e2' →
|
||||
base_step e1' e2' → contextual_step e1 e2.
|
||||
|
||||
Definition reducible (e : expr) :=
|
||||
∃ e', contextual_step e e'.
|
||||
|
||||
|
||||
Definition empty_ectx := HoleCtx.
|
||||
|
||||
(** Basic properties about the language *)
|
||||
Lemma fill_empty e : fill empty_ectx e = e.
|
||||
Proof. done. Qed.
|
||||
|
||||
Lemma fill_comp (K1 K2 : ectx) e : fill K1 (fill K2 e) = fill (comp_ectx K1 K2) e.
|
||||
Proof. induction K1; simpl; congruence. Qed.
|
||||
|
||||
|
||||
|
||||
Lemma base_contextual_step e1 e2 :
|
||||
base_step e1 e2 → contextual_step e1 e2.
|
||||
Proof. apply Ectx_step with empty_ectx; by rewrite ?fill_empty. Qed.
|
||||
|
||||
Lemma fill_contextual_step K e1 e2 :
|
||||
contextual_step e1 e2 → contextual_step (fill K e1) (fill K e2).
|
||||
Proof.
|
||||
destruct 1 as [K' e1' e2' -> ->].
|
||||
rewrite !fill_comp. by econstructor.
|
||||
Qed.
|
||||
|
||||
Fixpoint is_closed (X : list string) (e : expr) : bool :=
|
||||
match e with
|
||||
| Var x => bool_decide (x ∈ X)
|
||||
| Lam x e => is_closed (x :b: X) e
|
||||
| Unpack x e1 e2 => is_closed X e1 && is_closed (x :b: X) e2
|
||||
| Lit _ => true
|
||||
| UnOp _ e | Fst e | Snd e | InjL e | InjR e | TApp e | TLam e | Pack e => is_closed X e
|
||||
| App e1 e2 | BinOp _ e1 e2 | Pair e1 e2 => is_closed X e1 && is_closed X e2
|
||||
| If e0 e1 e2 | Case e0 e1 e2 =>
|
||||
is_closed X e0 && is_closed X e1 && is_closed X e2
|
||||
end.
|
||||
|
||||
(** [closed] states closedness as a Coq proposition, through the [Is_true] transformer. *)
|
||||
Definition closed (X : list string) (e : expr) : Prop := Is_true (is_closed X e).
|
||||
#[export] Instance closed_proof_irrel X e : ProofIrrel (closed X e).
|
||||
Proof. unfold closed. apply _. Qed.
|
||||
#[export] Instance closed_dec X e : Decision (closed X e).
|
||||
Proof. unfold closed. apply _. Defined.
|
||||
|
||||
(** closed expressions *)
|
||||
Lemma is_closed_weaken X Y e : is_closed X e → X ⊆ Y → is_closed Y e.
|
||||
Proof. revert X Y; induction e; naive_solver (eauto; set_solver). Qed.
|
||||
|
||||
Lemma is_closed_weaken_nil X e : is_closed [] e → is_closed X e.
|
||||
Proof. intros. by apply is_closed_weaken with [], list_subseteq_nil. Qed.
|
||||
|
||||
Lemma is_closed_subst X e x es :
|
||||
is_closed [] es → is_closed (x :: X) e → is_closed X (subst x es e).
|
||||
Proof.
|
||||
intros ?.
|
||||
induction e in X |-*; simpl; intros ?; destruct_and?; split_and?; simplify_option_eq;
|
||||
try match goal with
|
||||
| H : ¬(_ ∧ _) |- _ => apply not_and_l in H as [?%dec_stable|?%dec_stable]
|
||||
end; eauto using is_closed_weaken with set_solver.
|
||||
Qed.
|
||||
Lemma is_closed_do_subst' X e x es :
|
||||
is_closed [] es → is_closed (x :b: X) e → is_closed X (subst' x es e).
|
||||
Proof. destruct x; eauto using is_closed_subst. Qed.
|
||||
|
||||
(** Substitution lemmas *)
|
||||
Lemma subst_is_closed X e x es : is_closed X e → x ∉ X → subst x es e = e.
|
||||
Proof.
|
||||
induction e in X |-*; simpl; rewrite ?bool_decide_spec, ?andb_True; intros ??;
|
||||
repeat case_decide; simplify_eq; simpl; f_equal; intuition eauto with set_solver.
|
||||
Qed.
|
||||
|
||||
Lemma subst_is_closed_nil e x es : is_closed [] e → subst x es e = e.
|
||||
Proof. intros. apply subst_is_closed with []; set_solver. Qed.
|
||||
Lemma subst'_is_closed_nil e x es : is_closed [] e → subst' x es e = e.
|
||||
Proof. intros. destruct x as [ | x]. { done. } by apply subst_is_closed_nil. Qed.
|
||||
|
||||
|
||||
(* we derive a few lemmas about contextual steps *)
|
||||
Lemma contextual_step_app_l e1 e1' e2:
|
||||
is_val e2 →
|
||||
contextual_step e1 e1' →
|
||||
contextual_step (App e1 e2) (App e1' e2).
|
||||
Proof.
|
||||
intros [v <-%of_to_val]%is_val_spec Hcontextual.
|
||||
by eapply (fill_contextual_step (AppLCtx HoleCtx v)).
|
||||
Qed.
|
||||
|
||||
Lemma contextual_step_app_r e1 e2 e2':
|
||||
contextual_step e2 e2' →
|
||||
contextual_step (App e1 e2) (App e1 e2').
|
||||
Proof.
|
||||
intros Hcontextual.
|
||||
by eapply (fill_contextual_step (AppRCtx e1 HoleCtx)).
|
||||
Qed.
|
||||
|
||||
Lemma contextual_step_tapp e e':
|
||||
contextual_step e e' →
|
||||
contextual_step (TApp e) (TApp e').
|
||||
Proof.
|
||||
intros Hcontextual.
|
||||
by eapply (fill_contextual_step (TAppCtx HoleCtx)).
|
||||
Qed.
|
||||
|
||||
Lemma contextual_step_pack e e':
|
||||
contextual_step e e' →
|
||||
contextual_step (Pack e) (Pack e').
|
||||
Proof.
|
||||
intros Hcontextual.
|
||||
by eapply (fill_contextual_step (PackCtx HoleCtx)).
|
||||
Qed.
|
||||
|
||||
Lemma contextual_step_unpack x e e' e2:
|
||||
contextual_step e e' →
|
||||
contextual_step (Unpack x e e2) (Unpack x e' e2).
|
||||
Proof.
|
||||
intros Hcontextual.
|
||||
by eapply (fill_contextual_step (UnpackCtx x HoleCtx e2)).
|
||||
Qed.
|
||||
|
||||
Lemma contextual_step_unop op e e':
|
||||
contextual_step e e' →
|
||||
contextual_step (UnOp op e) (UnOp op e').
|
||||
Proof.
|
||||
intros Hcontextual.
|
||||
by eapply (fill_contextual_step (UnOpCtx op HoleCtx)).
|
||||
Qed.
|
||||
|
||||
Lemma contextual_step_binop_l op e1 e1' e2:
|
||||
is_val e2 →
|
||||
contextual_step e1 e1' →
|
||||
contextual_step (BinOp op e1 e2) (BinOp op e1' e2).
|
||||
Proof.
|
||||
intros [v <-%of_to_val]%is_val_spec Hcontextual.
|
||||
by eapply (fill_contextual_step (BinOpLCtx op HoleCtx v)).
|
||||
Qed.
|
||||
|
||||
Lemma contextual_step_binop_r op e1 e2 e2':
|
||||
contextual_step e2 e2' →
|
||||
contextual_step (BinOp op e1 e2) (BinOp op e1 e2').
|
||||
Proof.
|
||||
intros Hcontextual.
|
||||
by eapply (fill_contextual_step (BinOpRCtx op e1 HoleCtx)).
|
||||
Qed.
|
||||
|
||||
Lemma contextual_step_if e e' e1 e2:
|
||||
contextual_step e e' →
|
||||
contextual_step (If e e1 e2) (If e' e1 e2).
|
||||
Proof.
|
||||
intros Hcontextual.
|
||||
by eapply (fill_contextual_step (IfCtx HoleCtx e1 e2)).
|
||||
Qed.
|
||||
|
||||
Lemma contextual_step_pair_l e1 e1' e2:
|
||||
is_val e2 →
|
||||
contextual_step e1 e1' →
|
||||
contextual_step (Pair e1 e2) (Pair e1' e2).
|
||||
Proof.
|
||||
intros [v <-%of_to_val]%is_val_spec Hcontextual.
|
||||
by eapply (fill_contextual_step (PairLCtx HoleCtx v)).
|
||||
Qed.
|
||||
|
||||
Lemma contextual_step_pair_r e1 e2 e2':
|
||||
contextual_step e2 e2' →
|
||||
contextual_step (Pair e1 e2) (Pair e1 e2').
|
||||
Proof.
|
||||
intros Hcontextual.
|
||||
by eapply (fill_contextual_step (PairRCtx e1 HoleCtx)).
|
||||
Qed.
|
||||
|
||||
|
||||
Lemma contextual_step_fst e e':
|
||||
contextual_step e e' →
|
||||
contextual_step (Fst e) (Fst e').
|
||||
Proof.
|
||||
intros Hcontextual.
|
||||
by eapply (fill_contextual_step (FstCtx HoleCtx)).
|
||||
Qed.
|
||||
|
||||
Lemma contextual_step_snd e e':
|
||||
contextual_step e e' →
|
||||
contextual_step (Snd e) (Snd e').
|
||||
Proof.
|
||||
intros Hcontextual.
|
||||
by eapply (fill_contextual_step (SndCtx HoleCtx)).
|
||||
Qed.
|
||||
|
||||
Lemma contextual_step_injl e e':
|
||||
contextual_step e e' →
|
||||
contextual_step (InjL e) (InjL e').
|
||||
Proof.
|
||||
intros Hcontextual.
|
||||
by eapply (fill_contextual_step (InjLCtx HoleCtx)).
|
||||
Qed.
|
||||
|
||||
Lemma contextual_step_injr e e':
|
||||
contextual_step e e' →
|
||||
contextual_step (InjR e) (InjR e').
|
||||
Proof.
|
||||
intros Hcontextual.
|
||||
by eapply (fill_contextual_step (InjRCtx HoleCtx)).
|
||||
Qed.
|
||||
|
||||
Lemma contextual_step_case e e' e1 e2:
|
||||
contextual_step e e' →
|
||||
contextual_step (Case e e1 e2) (Case e' e1 e2).
|
||||
Proof.
|
||||
intros Hcontextual.
|
||||
by eapply (fill_contextual_step (CaseCtx HoleCtx e1 e2)).
|
||||
Qed.
|
||||
|
||||
#[export]
|
||||
Hint Resolve
|
||||
contextual_step_app_l contextual_step_app_r contextual_step_binop_l contextual_step_binop_r
|
||||
contextual_step_case contextual_step_fst contextual_step_if contextual_step_injl contextual_step_injr
|
||||
contextual_step_pack contextual_step_pair_l contextual_step_pair_r contextual_step_snd contextual_step_tapp
|
||||
contextual_step_tapp contextual_step_unop contextual_step_unpack : core.
|
@ -0,0 +1,90 @@
|
||||
From semantics.ts.systemf Require Export lang.
|
||||
Set Default Proof Using "Type".
|
||||
|
||||
|
||||
(** Coercions to make programs easier to type. *)
|
||||
Coercion of_val : val >-> expr.
|
||||
Coercion LitInt : Z >-> base_lit.
|
||||
Coercion LitBool : bool >-> base_lit.
|
||||
Coercion App : expr >-> Funclass.
|
||||
Coercion Var : string >-> expr.
|
||||
|
||||
(** Define some derived forms. *)
|
||||
Notation Let x e1 e2 := (App (Lam x e2) e1) (only parsing).
|
||||
Notation Seq e1 e2 := (Let BAnon e1 e2) (only parsing).
|
||||
Notation Match e0 x1 e1 x2 e2 := (Case e0 (Lam x1 e1) (Lam x2 e2)) (only parsing).
|
||||
|
||||
(* No scope for the values, does not conflict and scope is often not inferred
|
||||
properly. *)
|
||||
Notation "# l" := (LitV l%Z%V%stdpp) (at level 8, format "# l").
|
||||
Notation "# l" := (Lit l%Z%E%stdpp) (at level 8, format "# l") : expr_scope.
|
||||
|
||||
(** Syntax inspired by Coq/Ocaml. Constructions with higher precedence come
|
||||
first. *)
|
||||
Notation "( e1 , e2 , .. , en )" := (Pair .. (Pair e1 e2) .. en) : expr_scope.
|
||||
Notation "( e1 , e2 , .. , en )" := (PairV .. (PairV e1 e2) .. en) : val_scope.
|
||||
|
||||
Notation "'match:' e0 'with' 'InjL' x1 => e1 | 'InjR' x2 => e2 'end'" :=
|
||||
(Match e0 x1%binder e1 x2%binder e2)
|
||||
(e0, x1, e1, x2, e2 at level 200,
|
||||
format "'[hv' 'match:' e0 'with' '/ ' '[' 'InjL' x1 => '/ ' e1 ']' '/' '[' | 'InjR' x2 => '/ ' e2 ']' '/' 'end' ']'") : expr_scope.
|
||||
Notation "'match:' e0 'with' 'InjR' x1 => e1 | 'InjL' x2 => e2 'end'" :=
|
||||
(Match e0 x2%binder e2 x1%binder e1)
|
||||
(e0, x1, e1, x2, e2 at level 200, only parsing) : expr_scope.
|
||||
|
||||
Notation "()" := LitUnit : val_scope.
|
||||
|
||||
Notation "e1 + e2" := (BinOp PlusOp e1%E e2%E) : expr_scope.
|
||||
Notation "e1 - e2" := (BinOp MinusOp e1%E e2%E) : expr_scope.
|
||||
Notation "e1 * e2" := (BinOp MultOp e1%E e2%E) : expr_scope.
|
||||
Notation "e1 ≤ e2" := (BinOp LeOp e1%E e2%E) : expr_scope.
|
||||
Notation "e1 = e2" := (BinOp EqOp e1%E e2%E) : expr_scope.
|
||||
Notation "e1 < e2" := (BinOp LtOp e1%E e2%E) : expr_scope.
|
||||
|
||||
(*Notation "~ e" := (UnOp NegOp e%E) (at level 75, right associativity) : expr_scope.*)
|
||||
|
||||
Notation "'if:' e1 'then' e2 'else' e3" := (If e1%E e2%E e3%E)
|
||||
(at level 200, e1, e2, e3 at level 200) : expr_scope.
|
||||
|
||||
Notation "λ: x , e" := (Lam x%binder e%E)
|
||||
(at level 200, x at level 1, e at level 200,
|
||||
format "'[' 'λ:' x , '/ ' e ']'") : expr_scope.
|
||||
Notation "λ: x y .. z , e" := (Lam x%binder (Lam y%binder .. (Lam z%binder e%E) ..))
|
||||
(at level 200, x, y, z at level 1, e at level 200,
|
||||
format "'[' 'λ:' x y .. z , '/ ' e ']'") : expr_scope.
|
||||
|
||||
Notation "λ: x , e" := (LamV x%binder e%E)
|
||||
(at level 200, x at level 1, e at level 200,
|
||||
format "'[' 'λ:' x , '/ ' e ']'") : val_scope.
|
||||
Notation "λ: x y .. z , e" := (LamV x%binder (Lam y%binder .. (Lam z%binder e%E) .. ))
|
||||
(at level 200, x, y, z at level 1, e at level 200,
|
||||
format "'[' 'λ:' x y .. z , '/ ' e ']'") : val_scope.
|
||||
|
||||
Notation "'let:' x := e1 'in' e2" := (Lam x%binder e2%E e1%E)
|
||||
(at level 200, x at level 1, e1, e2 at level 200,
|
||||
format "'[' 'let:' x := '[' e1 ']' 'in' '/' e2 ']'") : expr_scope.
|
||||
Notation "e1 ;; e2" := (Lam BAnon e2%E e1%E)
|
||||
(at level 100, e2 at level 200,
|
||||
format "'[' '[hv' '[' e1 ']' ;; ']' '/' e2 ']'") : expr_scope.
|
||||
|
||||
|
||||
Notation "'Λ' , e" := (TLam e%E)
|
||||
(at level 200, e at level 200,
|
||||
format "'[' 'Λ' , '/ ' e ']'") : expr_scope.
|
||||
Notation "'Λ' , e" := (TLamV e%E)
|
||||
(at level 200, e at level 200,
|
||||
format "'[' 'Λ' , '/ ' e ']'") : val_scope.
|
||||
|
||||
(* the [e] always needs to be paranthesized, due to the level
|
||||
(chosen to make this cooperate with the [App] coercion) *)
|
||||
Notation "e '<>'" := (TApp e%E)
|
||||
(at level 10) : expr_scope.
|
||||
(*Check ((Λ, #4) <>)%E.*)
|
||||
(*Check (((λ: "x", "x") #5) <>)%E.*)
|
||||
|
||||
Notation "'pack' e" := (Pack e%E)
|
||||
(at level 200, e at level 200) : expr_scope.
|
||||
Notation "'pack' v" := (PackV v%V)
|
||||
(at level 200, v at level 200) : val_scope.
|
||||
Notation "'unpack' e1 'as' x 'in' e2" := (Unpack x%binder e1%E e2%E)
|
||||
(at level 200, e1, e2 at level 200, x at level 1) : expr_scope.
|
@ -0,0 +1,180 @@
|
||||
From stdpp Require Import prelude.
|
||||
From iris Require Import prelude.
|
||||
From semantics.ts.systemf Require Import lang.
|
||||
From semantics.lib Require Import maps.
|
||||
|
||||
Fixpoint subst_map (xs : gmap string expr) (e : expr) : expr :=
|
||||
match e with
|
||||
| Lit l => Lit l
|
||||
| Var y => match xs !! y with Some es => es | _ => Var y end
|
||||
| App e1 e2 => App (subst_map xs e1) (subst_map xs e2)
|
||||
| Lam x e => Lam x (subst_map (binder_delete x xs) e)
|
||||
| UnOp op e => UnOp op (subst_map xs e)
|
||||
| BinOp op e1 e2 => BinOp op (subst_map xs e1) (subst_map xs e2)
|
||||
| If e0 e1 e2 => If (subst_map xs e0) (subst_map xs e1) (subst_map xs e2)
|
||||
| TApp e => TApp (subst_map xs e)
|
||||
| TLam e => TLam (subst_map xs e)
|
||||
| Pack e => Pack (subst_map xs e)
|
||||
| Unpack x e1 e2 => Unpack x (subst_map xs e1) (subst_map (binder_delete x xs) e2)
|
||||
| Pair e1 e2 => Pair (subst_map xs e1) (subst_map xs e2)
|
||||
| Fst e => Fst (subst_map xs e)
|
||||
| Snd e => Snd (subst_map xs e)
|
||||
| InjL e => InjL (subst_map xs e)
|
||||
| InjR e => InjR (subst_map xs e)
|
||||
| Case e e1 e2 => Case (subst_map xs e) (subst_map xs e1) (subst_map xs e2)
|
||||
end.
|
||||
|
||||
Lemma subst_map_empty e :
|
||||
subst_map ∅ e = e.
|
||||
Proof.
|
||||
induction e; simpl; f_equal; eauto.
|
||||
all: destruct x; simpl; [done | by rewrite !delete_empty..].
|
||||
Qed.
|
||||
|
||||
Lemma subst_map_is_closed X e xs :
|
||||
is_closed X e →
|
||||
(∀ x : string, x ∈ dom xs → x ∉ X) →
|
||||
subst_map xs e = e.
|
||||
Proof.
|
||||
intros Hclosed Hd.
|
||||
induction e in xs, X, Hd, Hclosed |-*; simpl in *;try done.
|
||||
all: repeat match goal with
|
||||
| H : Is_true (_ && _) |- _ => apply andb_True in H as [ ? ? ]
|
||||
end.
|
||||
{ (* Var *)
|
||||
apply bool_decide_spec in Hclosed.
|
||||
assert (xs !! x = None) as ->; last done.
|
||||
destruct (xs !! x) as [s | ] eqn:Helem; last done.
|
||||
exfalso; eapply Hd; last apply Hclosed.
|
||||
apply elem_of_dom; eauto.
|
||||
}
|
||||
{ (* lambdas *)
|
||||
erewrite IHe; [done | done |].
|
||||
intros y. destruct x as [ | x]; first apply Hd.
|
||||
simpl.
|
||||
rewrite dom_delete elem_of_difference not_elem_of_singleton.
|
||||
intros [Hnx%Hd Hneq]. rewrite elem_of_cons. intros [? | ?]; done.
|
||||
}
|
||||
8: { (* Unpack *)
|
||||
erewrite IHe1; [ | done | done].
|
||||
erewrite IHe2; [ done | done | ].
|
||||
intros y. destruct x as [ | x]; first apply Hd.
|
||||
simpl.
|
||||
rewrite dom_delete elem_of_difference not_elem_of_singleton.
|
||||
intros [Hnx%Hd Hneq]. rewrite elem_of_cons. intros [? | ?]; done.
|
||||
}
|
||||
(* all other cases *)
|
||||
all: repeat match goal with
|
||||
| H : ∀ _ _, _ → _ → subst_map _ _ = _ |- _ => erewrite H; clear H
|
||||
end; done.
|
||||
Qed.
|
||||
|
||||
Lemma subst_map_subst map x (e e' : expr) :
|
||||
is_closed [] e' →
|
||||
subst_map map (subst x e' e) = subst_map (<[x:=e']>map) e.
|
||||
Proof.
|
||||
intros He'.
|
||||
revert x map; induction e; intros xx map; simpl;
|
||||
try (f_equal; eauto).
|
||||
- case_decide.
|
||||
+ simplify_eq/=. rewrite lookup_insert.
|
||||
rewrite (subst_map_is_closed []); [done | apply He' | ].
|
||||
intros ? ?. apply not_elem_of_nil.
|
||||
+ rewrite lookup_insert_ne; done.
|
||||
- destruct x; simpl; first done.
|
||||
+ case_decide.
|
||||
* simplify_eq/=. by rewrite delete_insert_delete.
|
||||
* rewrite delete_insert_ne; last by congruence. done.
|
||||
- destruct x; simpl; first done.
|
||||
+ case_decide.
|
||||
* simplify_eq/=. by rewrite delete_insert_delete.
|
||||
* rewrite delete_insert_ne; last by congruence. done.
|
||||
Qed.
|
||||
|
||||
Definition subst_is_closed (X : list string) (map : gmap string expr) :=
|
||||
∀ x e, map !! x = Some e → closed X e.
|
||||
Lemma subst_is_closed_subseteq X map1 map2 :
|
||||
map1 ⊆ map2 → subst_is_closed X map2 → subst_is_closed X map1.
|
||||
Proof.
|
||||
intros Hsub Hclosed2 x e Hl. eapply Hclosed2, map_subseteq_spec; done.
|
||||
Qed.
|
||||
|
||||
Lemma subst_subst_map x es map e :
|
||||
subst_is_closed [] map →
|
||||
subst x es (subst_map (delete x map) e) =
|
||||
subst_map (<[x:=es]>map) e.
|
||||
Proof.
|
||||
revert map es x; induction e; intros map v0 xx Hclosed; simpl;
|
||||
try (f_equal; eauto).
|
||||
- destruct (decide (xx=x)) as [->|Hne].
|
||||
+ rewrite lookup_delete // lookup_insert //. simpl.
|
||||
rewrite decide_True //.
|
||||
+ rewrite lookup_delete_ne // lookup_insert_ne //.
|
||||
destruct (_ !! x) as [rr|] eqn:Helem.
|
||||
* apply Hclosed in Helem.
|
||||
by apply subst_is_closed_nil.
|
||||
* simpl. rewrite decide_False //.
|
||||
- destruct x; simpl; first by auto.
|
||||
case_decide.
|
||||
+ simplify_eq. rewrite delete_idemp delete_insert_delete. done.
|
||||
+ rewrite delete_insert_ne //; last congruence. rewrite delete_commute. apply IHe.
|
||||
eapply subst_is_closed_subseteq; last done.
|
||||
apply map_delete_subseteq.
|
||||
- destruct x; simpl; first by auto.
|
||||
case_decide.
|
||||
+ simplify_eq. rewrite delete_idemp delete_insert_delete. done.
|
||||
+ rewrite delete_insert_ne //; last congruence. rewrite delete_commute. apply IHe2.
|
||||
eapply subst_is_closed_subseteq; last done.
|
||||
apply map_delete_subseteq.
|
||||
Qed.
|
||||
|
||||
Lemma subst'_subst_map b (es : expr) map e :
|
||||
subst_is_closed [] map →
|
||||
subst' b es (subst_map (binder_delete b map) e) =
|
||||
subst_map (binder_insert b es map) e.
|
||||
Proof.
|
||||
destruct b; first done.
|
||||
apply subst_subst_map.
|
||||
Qed.
|
||||
|
||||
Lemma closed_subst_weaken e map X Y :
|
||||
subst_is_closed [] map →
|
||||
(∀ x, x ∈ X → x ∉ dom map → x ∈ Y) →
|
||||
closed X e →
|
||||
closed Y (subst_map map e).
|
||||
Proof.
|
||||
induction e in X, Y, map |-*; rewrite /closed; simpl; intros Hmclosed Hsub Hcl; first done.
|
||||
all: repeat match goal with
|
||||
| H : Is_true (_ && _) |- _ => apply andb_True in H as [ ? ? ]
|
||||
end.
|
||||
{ (* vars *)
|
||||
destruct (map !! x) as [es | ] eqn:Heq.
|
||||
+ apply is_closed_weaken_nil. by eapply Hmclosed.
|
||||
+ apply bool_decide_pack. apply Hsub; first by eapply bool_decide_unpack.
|
||||
by apply not_elem_of_dom.
|
||||
}
|
||||
{ (* lambdas *)
|
||||
eapply IHe; last done.
|
||||
+ eapply subst_is_closed_subseteq; last done.
|
||||
destruct x; first done. apply map_delete_subseteq.
|
||||
+ intros y. destruct x as [ | x]; first by apply Hsub.
|
||||
rewrite !elem_of_cons. intros [-> | Hy] Hn; first by left.
|
||||
destruct (decide (y = x)) as [ -> | Hneq]; first by left.
|
||||
right. eapply Hsub; first done. set_solver.
|
||||
}
|
||||
8: { (* unpack *)
|
||||
apply andb_True; split; first naive_solver.
|
||||
eapply IHe2; last done.
|
||||
+ eapply subst_is_closed_subseteq; last done.
|
||||
destruct x; first done. apply map_delete_subseteq.
|
||||
+ intros y. destruct x as [ | x]; first by apply Hsub.
|
||||
rewrite !elem_of_cons. intros [-> | Hy] Hn; first by left.
|
||||
destruct (decide (y = x)) as [ -> | Hneq]; first by left.
|
||||
right. eapply Hsub; first done. set_solver.
|
||||
}
|
||||
(* all other cases *)
|
||||
all: repeat match goal with
|
||||
| |- Is_true (_ && _) => apply andb_True; split
|
||||
end.
|
||||
all: naive_solver.
|
||||
Qed.
|
@ -0,0 +1,113 @@
|
||||
From stdpp Require Import base relations.
|
||||
From iris Require Import prelude.
|
||||
From semantics.lib Require Import facts.
|
||||
From semantics.ts.systemf Require Export lang notation parallel_subst types bigstep.
|
||||
|
||||
Lemma list_subseteq_cons {X} (A B : list X) x : A ⊆ B → x :: A ⊆ x :: B.
|
||||
Proof. intros Hincl. intros y. rewrite !elem_of_cons. naive_solver. Qed.
|
||||
Lemma list_subseteq_cons_binder A B x : A ⊆ B → x :b: A ⊆ x :b: B.
|
||||
Proof. destruct x; first done. apply list_subseteq_cons. Qed.
|
||||
|
||||
Ltac simplify_list_elem :=
|
||||
simpl;
|
||||
repeat match goal with
|
||||
| |- ?x ∈ ?y :: ?l => apply elem_of_cons; first [left; reflexivity | right]
|
||||
| |- _ ∉ [] => apply not_elem_of_nil
|
||||
| |- _ ∉ _ :: _ => apply not_elem_of_cons; split
|
||||
end; try fast_done.
|
||||
Ltac simplify_list_subseteq :=
|
||||
simpl;
|
||||
repeat match goal with
|
||||
| |- ?x :: _ ⊆ ?x :: _ => apply list_subseteq_cons_l
|
||||
| |- ?x :: _ ⊆ _ => apply list_subseteq_cons_elem; first solve [simplify_list_elem]
|
||||
| |- elements _ ⊆ elements _ => apply elements_subseteq; set_solver
|
||||
| |- [] ⊆ _ => apply list_subseteq_nil
|
||||
| |- ?x :b: _ ⊆ ?x :b: _ => apply list_subseteq_cons_binder
|
||||
end;
|
||||
try fast_done.
|
||||
|
||||
(* Try to solve [is_closed] goals using a number of heuristics (that shouldn't make the goal unprovable) *)
|
||||
Ltac simplify_closed :=
|
||||
simpl; intros;
|
||||
repeat match goal with
|
||||
| |- closed _ _ => unfold closed; simpl
|
||||
| |- Is_true (is_closed [] _) => first [assumption | done]
|
||||
| |- Is_true (is_closed _ (lang.subst _ _ _)) => rewrite subst_is_closed_nil; last solve[simplify_closed]
|
||||
| |- Is_true (is_closed ?X ?v) => assert_fails (is_evar X); eapply is_closed_weaken
|
||||
| |- Is_true (is_closed _ _) => eassumption
|
||||
| |- Is_true (_ && true) => rewrite andb_true_r
|
||||
| |- Is_true (true && _) => rewrite andb_true_l
|
||||
| |- Is_true (?a && ?a) => rewrite andb_diag
|
||||
| |- Is_true (_ && _) => simpl; rewrite !andb_True; split_and!
|
||||
| |- _ ⊆ ?A => match type of A with | list _ => simplify_list_subseteq end
|
||||
| |- _ ∈ ?A => match type of A with | list _ => simplify_list_elem end
|
||||
| |- _ ∉ ?A => match type of A with | list _ => simplify_list_elem end
|
||||
| |- _ ≠ _ => congruence
|
||||
| H : closed _ _ |- _ => unfold closed in H; simpl in H
|
||||
| H: Is_true (_ && _) |- _ => simpl in H; apply andb_True in H
|
||||
| H : _ ∧ _ |- _ => destruct H
|
||||
| |- Is_true (bool_decide (_ ∈ _)) => apply bool_decide_pack; set_solver
|
||||
end; try fast_done.
|
||||
|
||||
|
||||
Ltac bs_step_det :=
|
||||
repeat match goal with
|
||||
| |- big_step ?e _ =>
|
||||
simpl;
|
||||
match e with
|
||||
(* case analysis, don't backtrack but pose the goal to the user *)
|
||||
| Case _ _ _ => idtac
|
||||
| If _ _ _ => idtac
|
||||
(* use the value lemma *)
|
||||
| of_val _ => apply big_step_of_val; done
|
||||
(* try to solve substitutions by using disjointness assumptions *)
|
||||
| context[decide (?p = ?p)] => rewrite decide_True; [ | done]
|
||||
| context[decide (_ = _)] => rewrite decide_True; [ | congruence]
|
||||
| context[decide (_ = _)] => rewrite decide_False; [ | congruence]
|
||||
| context[decide (_ ≠ _)] => rewrite decide_True; [ | congruence]
|
||||
| context[decide (_ ≠ _)] => rewrite decide_False; [ | congruence]
|
||||
(* if we have a substitution that didn't simplify, try to show that it's closed *)
|
||||
(* we don't make any attempt to solve it if it isn't closed under [] *)
|
||||
| context[lang.subst _ _ ?e] => is_var e; rewrite subst_is_closed_nil; last solve[simplify_closed]
|
||||
| context[lang.subst _ _ (of_val ?v)] => is_var v; rewrite subst_is_closed_nil; last solve[simplify_closed]
|
||||
(* try to use a [big_step] assumption, or else apply a constructor *)
|
||||
| _ => first [eassumption | econstructor]
|
||||
end
|
||||
| |- bin_op_eval _ _ _ = _ =>
|
||||
simpl; reflexivity
|
||||
end; simplify_closed.
|
||||
Ltac bs_steps_det := repeat bs_step_det.
|
||||
|
||||
Ltac map_solver :=
|
||||
repeat match goal with
|
||||
| |- (⤉ _) !! _ = Some _ =>
|
||||
rewrite fmap_insert
|
||||
| |- <[ ?p := _ ]> _ !! ?p = Some _ =>
|
||||
apply lookup_insert
|
||||
| |- <[ _ := _ ]> _ !! _ = Some _ =>
|
||||
rewrite lookup_insert_ne; [ | congruence]
|
||||
end.
|
||||
Ltac solve_typing :=
|
||||
repeat match goal with
|
||||
| |- TY _ ; _ ⊢ ?e : ?A => first [eassumption | econstructor]
|
||||
(* heuristic to solve tapp goals where we need to pick the right type for the substitution *)
|
||||
| |- TY _ ; _ ⊢ ?e <> : ?A => eapply typed_tapp'; [solve_typing | | asimpl; reflexivity]
|
||||
| |- bin_op_typed _ _ _ _ => econstructor
|
||||
| |- type_wf _ ?e => assert_fails (is_evar e); eassumption
|
||||
| |- type_wf _ ?e => assert_fails (is_evar e); econstructor
|
||||
| |- type_wf _ (subst _ ?A) => eapply type_wf_subst; [ | intros; simpl]
|
||||
| |- type_wf _ ?e => assert_fails (is_evar e); eapply type_wf_mono; [eassumption | lia]
|
||||
(* conditions spawned by the tyvar case of [type_wf] *)
|
||||
| |- _ < _ => lia
|
||||
(* conditions spawned by the variable case *)
|
||||
| |- _ !! _ = Some _ => map_solver
|
||||
end.
|
||||
|
||||
Tactic Notation "unify_type" uconstr(A) :=
|
||||
match goal with
|
||||
| |- TY ?n; ?Γ ⊢ ?e : ?B => unify A B
|
||||
end.
|
||||
Tactic Notation "replace_type" uconstr(A) :=
|
||||
match goal with
|
||||
| |- TY ?n; ?Γ ⊢ ?e : ?B => replace B%ty with A%ty; cycle -1; first try by asimpl
|
||||
end.
|
Loading…
Reference in new issue