mirror of
https://github.com/achlipala/frap.git
synced 2024-11-10 00:07:51 +00:00
Finish port to Coq 8.6
This commit is contained in:
parent
1768aa6ea7
commit
466ea72b27
3 changed files with 15 additions and 6 deletions
7
Frap.v
7
Frap.v
|
@ -69,10 +69,13 @@ Ltac same_structure x y :=
|
|||
|
||||
Ltac instantiate_obvious1 H :=
|
||||
match type of H with
|
||||
| _ ++ _ = _ ++ _ -> _ => fail 1
|
||||
| ?x = ?y -> _ =>
|
||||
(same_structure x y; specialize (H eq_refl)) || fail 3
|
||||
(same_structure x y; specialize (H eq_refl))
|
||||
|| (has_evar (x, y); fail 3)
|
||||
| JMeq.JMeq ?x ?y -> _ =>
|
||||
(same_structure x y; specialize (H JMeq.JMeq_refl)) || fail 3
|
||||
(same_structure x y; specialize (H JMeq.JMeq_refl))
|
||||
|| (has_evar (x, y); fail 3)
|
||||
| forall x : ?T, _ =>
|
||||
match type of T with
|
||||
| Prop => fail 1
|
||||
|
|
|
@ -1496,6 +1496,7 @@ Proof.
|
|||
cases r1; simplify; subst.
|
||||
right; eexists (_ :: _); simplify; eauto.
|
||||
invert H.
|
||||
first_order; subst; eauto.
|
||||
apply IHl1 in H2; first_order; subst; eauto.
|
||||
Qed.
|
||||
|
||||
|
@ -1550,7 +1551,7 @@ Proof.
|
|||
|
||||
cases x3; simplify.
|
||||
invert H.
|
||||
eapply nextAction_det in H0; eauto; propositional; subst.
|
||||
eapply nextAction_det in H0; try eapply H5; eauto; propositional; subst.
|
||||
eauto 10.
|
||||
|
||||
invert H.
|
||||
|
@ -1574,7 +1575,7 @@ Proof.
|
|||
|
||||
cases x3; simplify.
|
||||
invert H.
|
||||
eapply nextAction_det in H0; eauto; propositional; subst.
|
||||
eapply nextAction_det in H0; try eapply H5; eauto; propositional; subst.
|
||||
eauto 10.
|
||||
|
||||
invert H.
|
||||
|
|
|
@ -236,6 +236,7 @@ Module References.
|
|||
| [ H : step _ _ |- _ ] => invert H
|
||||
end;
|
||||
repeat match goal with
|
||||
| [ H : plug _ _ _ |- _ ] => invert1 H
|
||||
| [ H : plug _ _ _ |- _ ] => invert H
|
||||
| [ H : step0 _ _ |- _ ] => invert1 H
|
||||
| [ H : value _ |- _ ] => invert1 H
|
||||
|
@ -302,8 +303,12 @@ Module References.
|
|||
\/ exists he', step (h, e) he'.
|
||||
Proof.
|
||||
induct 2; t.
|
||||
apply H2 in H8; t.
|
||||
apply H1 in H8; t.
|
||||
match goal with
|
||||
| [ H1 : _ = Some _, H2 : forall l : loc, _ |- _ ] => apply H3 in H8; t
|
||||
end.
|
||||
match goal with
|
||||
| [ H1 : _ = Some _, H2 : forall l : loc, _ |- _ ] => apply H3 in H8; t
|
||||
end.
|
||||
Qed.
|
||||
|
||||
(* Now, a series of lemmas essentially copied from original type-soundness
|
||||
|
|
Loading…
Reference in a new issue