mpri-funcprog-project/coq/LambdaCalculusFreeVars.v
2017-10-05 17:57:33 +02:00

127 lines
2.8 KiB
Coq

Require Import MyTactics.
Require Import LambdaCalculusSyntax.
(* This technical lemma states that the renaming [lift 1] is injective. *)
Lemma lift_inj_Var:
forall t x,
lift 1 t = Var (S x) <-> t = Var x.
Proof.
split; intros.
{ eauto using lift_inj. }
{ subst. eauto. }
Qed.
(* -------------------------------------------------------------------------- *)
(* The predicate [fv] is characterized by the following lemmas. *)
Lemma fv_Var_eq:
forall k x,
fv k (Var x) <-> x < k.
Proof.
unfold fv. asimpl. induction k; intros.
(* Base case. *)
{ asimpl. split; intros; false.
{ unfold ids, Ids_term in *. injections. omega. }
{ omega. }
}
(* Step. *)
{ destruct x; asimpl.
{ split; intros. { omega. } { reflexivity. } }
rewrite lift_inj_Var. rewrite IHk. omega. }
Qed.
Lemma fv_Lam_eq:
forall k t,
fv k (Lam t) <-> fv (S k) t.
Proof.
unfold fv. intros. asimpl. split; intros.
{ injections. eauto. }
{ unpack. congruence. }
Qed.
Lemma fv_App_eq:
forall k t1 t2,
fv k (App t1 t2) <-> fv k t1 /\ fv k t2.
Proof.
unfold fv. intros. asimpl. split; intros.
{ injections. eauto. }
{ unpack. congruence. }
Qed.
Lemma fv_Let_eq:
forall k t1 t2,
fv k (Let t1 t2) <-> fv k t1 /\ fv (S k) t2.
Proof.
unfold fv. intros. asimpl. split; intros.
{ injections. eauto. }
{ unpack. congruence. }
Qed.
Hint Rewrite fv_Var_eq fv_Lam_eq fv_App_eq fv_Let_eq : fv.
(* -------------------------------------------------------------------------- *)
(* The following lemmas allow decomposing a closedness hypothesis.
Because [closed] is not an inductive notion, there is no lemma
for [Lam] and for the right-hand side of [Let]. *)
Lemma closed_Var:
forall x,
~ closed (Var x).
Proof.
unfold closed; intros; fv. omega.
Qed.
Lemma closed_AppL:
forall t1 t2,
closed (App t1 t2) ->
closed t1.
Proof.
unfold closed; intros; fv. tauto.
Qed.
Lemma closed_AppR:
forall t1 t2,
closed (App t1 t2) ->
closed t2.
Proof.
unfold closed; intros; fv. tauto.
Qed.
Lemma closed_LetL:
forall t1 t2,
closed (Let t1 t2) ->
closed t1.
Proof.
unfold closed; intros; fv. tauto.
Qed.
Hint Resolve closed_Var closed_AppL closed_AppR closed_LetL : closed.
(* -------------------------------------------------------------------------- *)
(* If the free variables of the term [t] are below [k], then [t] is unaffected
by a substitution of the form [upn k sigma]. *)
Lemma fv_unaffected:
forall t k sigma,
fv k t ->
t.[upn k sigma] = t.
Proof.
induction t; intros; fv; unpack; asimpl;
try solve [ eauto using upn_k_sigma_x with typeclass_instances
| f_equal; eauto ].
Qed.
(* If the term [t] is closed, then [t] is unaffected by any substitution. *)
Lemma closed_unaffected:
forall t sigma,
closed t ->
t.[sigma] = t.
Proof.
intros. eapply fv_unaffected with (k := 0). eauto.
Qed.