add expr_open_lc and expr_subst_open lemmas
This commit is contained in:
parent
080aa0ffec
commit
3d200e141e
1 changed files with 95 additions and 1 deletions
|
@ -1,3 +1,5 @@
|
|||
From Coq Require Import Lists.List.
|
||||
Import ListNotations.
|
||||
Require Import Atom.
|
||||
Require Import Metatheory.
|
||||
Require Import FSetNotin.
|
||||
|
@ -103,7 +105,7 @@ Lemma type_open_lc_core : forall (τ:type_DeBruijn) i (σ1:type_DeBruijn) j (σ2
|
|||
{i ~tt~> σ1} τ = {j ~tt~> σ2} ({i ~tt~> σ1} τ) ->
|
||||
({j ~tt~> σ2} τ) = τ
|
||||
.
|
||||
Proof with eauto*.
|
||||
Proof.
|
||||
induction τ;
|
||||
intros i σ1 j σ2 Neq H;
|
||||
simpl in *;
|
||||
|
@ -181,3 +183,95 @@ Proof.
|
|||
reflexivity.
|
||||
trivial.
|
||||
Qed.
|
||||
|
||||
|
||||
|
||||
|
||||
|
||||
|
||||
|
||||
|
||||
Lemma expr_open_lc_core : forall (t:expr_DeBruijn) i (s1:expr_DeBruijn) j (s2:expr_DeBruijn),
|
||||
i <> j ->
|
||||
{i ~ee~> s1} t = {j ~ee~> s2} ({i ~ee~> s1} t) ->
|
||||
({j ~ee~> s2} t) = t
|
||||
.
|
||||
Proof.
|
||||
induction t;
|
||||
intros i s1 j s2 Neq H;
|
||||
simpl in *;
|
||||
inversion H;
|
||||
f_equal; eauto.
|
||||
|
||||
(* bound var *)
|
||||
- simpl in *.
|
||||
destruct (j === n).
|
||||
destruct (i === n).
|
||||
3:reflexivity.
|
||||
|
||||
rewrite e,e0 in Neq.
|
||||
contradiction Neq.
|
||||
reflexivity.
|
||||
|
||||
rewrite H,e.
|
||||
simpl.
|
||||
destruct (n===n).
|
||||
reflexivity.
|
||||
contradict n1.
|
||||
reflexivity.
|
||||
Qed.
|
||||
|
||||
|
||||
(*
|
||||
* opening is idempotent on locally closed expressions
|
||||
*)
|
||||
Lemma expr_open_lc : forall k s t,
|
||||
expr_lc t ->
|
||||
({ k ~ee~> s } t) = t
|
||||
.
|
||||
Proof.
|
||||
intros.
|
||||
generalize dependent k.
|
||||
induction H; eauto; simpl in *; intro k; f_equal; eauto.
|
||||
|
||||
- unfold expr_open in *.
|
||||
pick fresh x for L.
|
||||
apply expr_open_lc_core with (i:=0) (s1:=(ex_fvar x)) (j:=S k) (s2:=s).
|
||||
discriminate.
|
||||
apply eq_sym, H1.
|
||||
assumption.
|
||||
|
||||
- unfold expr_open in *.
|
||||
pick fresh x for L.
|
||||
apply expr_open_lc_core with (i:=0) (s1:=(ex_fvar x)) (j:=S k) (s2:=s).
|
||||
discriminate.
|
||||
apply eq_sym, H1.
|
||||
assumption.
|
||||
Qed.
|
||||
|
||||
(*
|
||||
* type substitution distributes over opening
|
||||
*)
|
||||
Lemma expr_subst_open : forall t s1 s2 x k,
|
||||
expr_lc s2 ->
|
||||
|
||||
[x ~ee~> s2] ({k ~ee~> s1} t)
|
||||
=
|
||||
{k ~ee~> [x ~ee~> s2] s1} ([x ~ee~> s2] t).
|
||||
Proof.
|
||||
induction t;
|
||||
intros; simpl; f_equal; auto.
|
||||
|
||||
(* free var *)
|
||||
- destruct (x == a).
|
||||
subst.
|
||||
apply eq_sym, expr_open_lc.
|
||||
assumption.
|
||||
trivial.
|
||||
|
||||
(* bound var *)
|
||||
- destruct (k === n).
|
||||
reflexivity.
|
||||
trivial.
|
||||
Qed.
|
||||
|
||||
|
|
Loading…
Reference in a new issue