mirror of
https://github.com/achlipala/frap.git
synced 2024-11-12 17:17:50 +00:00
Ported to Coq 8.15
This commit is contained in:
parent
1b43aa9aea
commit
d745f0802e
11 changed files with 25 additions and 17 deletions
|
@ -426,7 +426,6 @@ Proof.
|
|||
simp.
|
||||
cancel; auto.
|
||||
subst; cancel.
|
||||
cancel; auto.
|
||||
Qed.
|
||||
|
||||
Theorem HtRead'' : forall linvs p P R,
|
||||
|
@ -448,6 +447,8 @@ Proof.
|
|||
constructor.
|
||||
simp.
|
||||
cancel.
|
||||
subst.
|
||||
assumption.
|
||||
Qed.
|
||||
|
||||
Ltac basic := apply HtReturn' || eapply HtWrite || eapply HtAlloc || eapply HtFree
|
||||
|
|
|
@ -407,7 +407,6 @@ Proof.
|
|||
simp.
|
||||
cancel; auto.
|
||||
subst; cancel.
|
||||
cancel; auto.
|
||||
Qed.
|
||||
|
||||
Theorem HtRead'' : forall linvs p P R,
|
||||
|
@ -429,6 +428,8 @@ Proof.
|
|||
constructor.
|
||||
simp.
|
||||
cancel.
|
||||
subst.
|
||||
assumption.
|
||||
Qed.
|
||||
|
||||
Ltac basic := apply HtReturn' || eapply HtWrite || eapply HtAlloc || eapply HtFree
|
||||
|
|
11
Connecting.v
11
Connecting.v
|
@ -582,6 +582,8 @@ Module MixedEmbedded(Import BW : BIT_WIDTH).
|
|||
constructor.
|
||||
simp.
|
||||
cancel.
|
||||
subst.
|
||||
assumption.
|
||||
Qed.
|
||||
|
||||
Transparent heq himp lift star exis ptsto.
|
||||
|
@ -651,7 +653,7 @@ Module MixedEmbedded(Import BW : BIT_WIDTH).
|
|||
repeat rewrite lookup_join2 by (simp; sets); reflexivity.
|
||||
unfold disjoint in *; simp.
|
||||
cases (weq a0 a); simp.
|
||||
apply H1 with (a0 := a).
|
||||
apply H1 with (a := a).
|
||||
unfold heap1; simp.
|
||||
equality.
|
||||
assumption.
|
||||
|
@ -817,7 +819,6 @@ Module MixedEmbedded(Import BW : BIT_WIDTH).
|
|||
simp.
|
||||
cancel; auto.
|
||||
subst; cancel.
|
||||
cancel; auto.
|
||||
Qed.
|
||||
|
||||
Theorem HtRead'' : forall p P R,
|
||||
|
@ -2666,7 +2667,7 @@ Module MixedToDeep(Import BW : BIT_WIDTH).
|
|||
rewrite lookup_join2 by (eapply lookup_None_dom; simplify; eauto).
|
||||
eassumption.
|
||||
|
||||
Grab Existential Variables.
|
||||
Unshelve.
|
||||
exact (^0) || exact (Return (^0)).
|
||||
exact (^0) || exact (Return (^0)).
|
||||
exact (^0) || exact (Return (^0)).
|
||||
|
@ -3159,7 +3160,7 @@ Module MixedToDeep(Import BW : BIT_WIDTH).
|
|||
eexists.
|
||||
repeat econstructor.
|
||||
|
||||
Grab Existential Variables.
|
||||
Unshelve.
|
||||
exact (^0) || exact (Return (^0)).
|
||||
exact (^0) || exact (Return (^0)).
|
||||
exact (^0) || exact (Return (^0)).
|
||||
|
@ -3313,8 +3314,6 @@ Module MixedToDeep(Import BW : BIT_WIDTH).
|
|||
simp.
|
||||
step.
|
||||
cancel.
|
||||
subst.
|
||||
cancel.
|
||||
Qed.
|
||||
|
||||
Theorem incrementer_compiled_ok : forall p v,
|
||||
|
|
|
@ -1328,7 +1328,7 @@ Module FindDuplicates (FS : FINITE_SET).
|
|||
|
||||
invert H0.
|
||||
exfalso.
|
||||
apply H with (x3 := x).
|
||||
apply H with (x := x).
|
||||
exists x0.
|
||||
exists x1.
|
||||
exists x2.
|
||||
|
|
|
@ -1144,7 +1144,7 @@ Module FindDuplicates (FS : FINITE_SET).
|
|||
|
||||
invert H0.
|
||||
exfalso.
|
||||
apply H with (x3 := x).
|
||||
apply H with (x := x).
|
||||
exists x0.
|
||||
exists x1.
|
||||
exists x2.
|
||||
|
|
|
@ -657,7 +657,7 @@ Module Deeper.
|
|||
end).
|
||||
cases (r ==n needle); ht.
|
||||
cases (i ==n acc); simp.
|
||||
apply H3 with (i0 := i); auto.
|
||||
apply H3 with (i := i); auto.
|
||||
Qed.
|
||||
|
||||
(* The single-stepping interpreter forms the basis for defining transition
|
||||
|
|
|
@ -454,3 +454,8 @@ Ltac unfold_recurse f k :=
|
|||
linear_arithmetic). *)
|
||||
Arguments N.mul: simpl never.
|
||||
Arguments N.add: simpl never.
|
||||
|
||||
Definition IF_then_else (p q1 q2 : Prop) :=
|
||||
(p /\ q1) \/ (~p /\ q2).
|
||||
|
||||
Notation "'IF' p 'then' q1 'else' q2" := (IF_then_else p q1 q2) (at level 95).
|
||||
|
|
|
@ -375,7 +375,7 @@ Proof.
|
|||
| [ H : (_, _) = (_, _) |- _ ] => invert H
|
||||
end; eauto.
|
||||
|
||||
Grab Existential Variables.
|
||||
Unshelve.
|
||||
exact 0.
|
||||
Qed.
|
||||
|
||||
|
|
|
@ -666,6 +666,8 @@ Proof.
|
|||
constructor.
|
||||
simp.
|
||||
cancel.
|
||||
subst.
|
||||
assumption.
|
||||
Qed.
|
||||
|
||||
(* Temporarily transparent again! *)
|
||||
|
@ -736,7 +738,7 @@ Proof.
|
|||
repeat rewrite lookup_join2 by (simp; sets); reflexivity.
|
||||
unfold disjoint in *; simp.
|
||||
cases (a0 ==n a); simp.
|
||||
apply H1 with (a0 := a).
|
||||
apply H1 with (a := a).
|
||||
unfold heap1; simp.
|
||||
equality.
|
||||
assumption.
|
||||
|
@ -959,7 +961,6 @@ Proof.
|
|||
simp.
|
||||
cancel; auto.
|
||||
subst; cancel.
|
||||
cancel; auto.
|
||||
Qed.
|
||||
|
||||
Theorem HtRead'' : forall p P R,
|
||||
|
|
|
@ -371,7 +371,6 @@ Proof.
|
|||
simp.
|
||||
cancel; auto.
|
||||
subst; cancel.
|
||||
cancel; auto.
|
||||
Qed.
|
||||
|
||||
Theorem HtRead'' : forall p P R,
|
||||
|
@ -393,6 +392,8 @@ Proof.
|
|||
constructor.
|
||||
simp.
|
||||
cancel.
|
||||
subst.
|
||||
assumption.
|
||||
Qed.
|
||||
|
||||
|
||||
|
@ -984,7 +985,7 @@ Proof.
|
|||
repeat rewrite lookup_join2 by (simp; sets); reflexivity.
|
||||
unfold disjoint in *; simp.
|
||||
cases (a0 ==n a); simp.
|
||||
apply H1 with (a0 := a).
|
||||
apply H1 with (a := a).
|
||||
unfold heap1; simp.
|
||||
equality.
|
||||
assumption.
|
||||
|
|
|
@ -6094,7 +6094,7 @@ The project home page is:
|
|||
\begin{center}
|
||||
\url{https://coq.inria.fr/}
|
||||
\end{center}
|
||||
The code associated with this book is designed to work with Coq versions 8.9 and higher.
|
||||
The code associated with this book is designed to work with Coq versions 8.11 and higher.
|
||||
The project Web site makes a number of versions available, and versions are also available in popular OS package distributions, along with binaries for platforms where open-source package systems are less common.
|
||||
We assume that readers have installed Coq by one of those means or another.
|
||||
It will also be almost essential to use some graphical interface for Coq editing.
|
||||
|
|
Loading…
Reference in a new issue