Commit 63f6560a authored by Adam Chlipala's avatar Adam Chlipala

ccTerm_correct

parent 402e0cca
...@@ -826,7 +826,7 @@ Section spliceFuncs_correct. ...@@ -826,7 +826,7 @@ Section spliceFuncs_correct.
End spliceFuncs_correct. End spliceFuncs_correct.
Notation "var <| to" := (match to with Notation "var <| to" := (match to return Set with
| None => unit | None => unit
| Some t => var (ccType t) | Some t => var (ccType t)
end) (no associativity, at level 70). end) (no associativity, at level 70).
...@@ -925,10 +925,10 @@ Section packing_correct. ...@@ -925,10 +925,10 @@ Section packing_correct.
Qed. Qed.
End packing_correct. End packing_correct.
(*Lemma typeDenote_same : forall t, About packExp_correct.
Closed.typeDenote (ccType t) = Source.typeDenote t. Implicit Arguments packExp_correct [v envT fvs1].
induction t; crush.
Qed.*) Implicit Arguments lookup_type_more [v2 envT fvs t b v].
Lemma typeDenote_same : forall t, Lemma typeDenote_same : forall t,
Source.typeDenote t = Closed.typeDenote (ccType t). Source.typeDenote t = Closed.typeDenote (ccType t).
...@@ -945,19 +945,13 @@ Qed. ...@@ -945,19 +945,13 @@ Qed.
Implicit Arguments look [envT n fvs t]. Implicit Arguments look [envT n fvs t].
Lemma cast_jmeq : forall (T1 T2 : Set) (pf : T1 = T2) (T2' : Set) (v1 : T1) (v2 : T2'), Fixpoint lr (t : Source.type) : Source.typeDenote t -> Closed.typeDenote (ccType t) -> Prop :=
v1 == v2 match t return Source.typeDenote t -> Closed.typeDenote (ccType t) -> Prop with
-> T2' = T2 | Nat => @eq nat
-> match pf in _ = T return T with | dom --> ran => fun f1 f2 =>
| refl_equal => v1 forall x1 x2, lr dom x1 x2
end == v2. -> lr ran (f1 x1) (f2 x2)
intros; generalize pf; subst. end.
intro.
rewrite (UIP_refl _ _ pf).
auto.
Qed.
Hint Resolve cast_jmeq.
Theorem ccTerm_correct : forall t G Theorem ccTerm_correct : forall t G
(e1 : Source.exp Source.typeDenote t) (e1 : Source.exp Source.typeDenote t)
...@@ -970,10 +964,15 @@ Theorem ccTerm_correct : forall t G ...@@ -970,10 +964,15 @@ Theorem ccTerm_correct : forall t G
-> v2 < length envT) -> v2 < length envT)
-> (forall t (v1 : Source.typeDenote t) (v2 : natvar t), -> (forall t (v1 : Source.typeDenote t) (v2 : natvar t),
In (existT _ _ (v1, v2)) G In (existT _ _ (v1, v2)) G
-> lookup_type v2 fvs = Some t -> forall pf,
-> lookup Closed.typeDenote v2 env == v1) lr t v1 (match lok pf in _ = T return T with
-> Closed.expDenote (funcsDenote (ccExp e2 fvs wf) env) | refl_equal => lookup Closed.typeDenote v2 env
== Source.expDenote e1. end))
(*-> forall pf : lookup_type v2 fvs = Some t,
lr t v1 (match pf in _ = T return Closed.typeDenote <| T with
| refl_equal => lookup Closed.typeDenote v2 env
end))*)
-> lr t (Source.expDenote e1) (Closed.expDenote (funcsDenote (ccExp e2 fvs wf) env)).
Hint Rewrite spliceFuncs_correct unpackExp_correct : cpdt. Hint Rewrite spliceFuncs_correct unpackExp_correct : cpdt.
Hint Resolve packExp_correct lookup_type_inner. Hint Resolve packExp_correct lookup_type_inner.
...@@ -985,37 +984,60 @@ Theorem ccTerm_correct : forall t G ...@@ -985,37 +984,60 @@ Theorem ccTerm_correct : forall t G
crush. crush.
crush. crush.
Lemma app_jmeq : forall dom1 dom2 ran1 ran2 crush.
(f1 : dom1 -> ran1) (f2 : dom2 -> ran2) apply (H0 x1 (length envT) (t1 :: envT) (true, snd (fvsExp (f2 (length envT)) (t1 :: envT)))).
(x1 : dom1) (x2 : dom2), clear H0.
ran1 = ran2 crush.
-> f1 == f2 eauto.
-> x1 == x2 generalize (H1 _ _ _ H4).
-> f1 x1 == f2 x2. crush.
crush.
assert (dom1 = dom2).
inversion H1; trivial.
crush.
Qed.
Lemma app_jmeq : forall dom ran
(f1 : Closed.typeDenote (ccType dom) -> Closed.typeDenote (ccType ran))
(f2 : Source.typeDenote dom -> Source.typeDenote ran)
(x1 : dom1) (x2 : dom2),
ran1 = ran2
-> f1 == f2
-> x1 == x2
-> f1 x1 == f2 x2.
crush.
assert (dom1 = dom2).
inversion H1; trivial.
crush.
Qed.
crush.
generalize H3; clear_all.
match goal with
| [ |- context[match ?pf with refl_equal => _ end] ] => generalize pf
end.
simpl. simpl.
refine (app_jmeq _ _ _). destruct (eq_nat_dec (length envT) (length envT)); crush.
apply app_jmeq. rewrite (UIP_refl _ _ e0); assumption.
dependent rewrite <- IHexp_equiv1.
match goal with
| [ |- context[match ?pf with refl_equal => _ end] ] => generalize pf
end; simpl.
destruct (eq_nat_dec v2 (length envT)); crush.
generalize (H1 _ _ _ H5).
crush.
generalize (H2 _ _ _ H5).
clear H2; intro H2.
generalize (H2 (lookup_type_inner _ _ _ _ _ wf pf)).
clear H2; intro H2.
generalize H2; clear_all.
repeat match goal with
| [ |- context[match ?pf with refl_equal => _ end] ] => generalize pf
end; simpl.
match type of pf with
| ?X = _ => assert (X <> None)
end.
congruence.
assert (forall (n : nat) (t0 : Source.type),
lookup_type (envT:=envT) n (snd (fvsExp (f2 (length envT)) (t1 :: envT))) = Some t0 ->
lookup_type (envT:=envT) n fvs = Some t0).
eauto.
generalize (packExp_correct fvs (inclusion t1 envT fvs (f2 (length envT)) wf) env H).
simpl.
match goal with
| [ |- ?X == ?Y -> _ ] =>
generalize X Y
end.
rewrite pf.
rewrite (lookup_type_inner _ _ _ _ _ wf pf).
crush.
rewrite (UIP_refl _ _ e0).
rewrite (UIP_refl _ _ e1) in H2.
crush.
rewrite <- H1.
assumption.
Qed. Qed.
......
Markdown is supported
0% or
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment