mirror of
https://github.com/achlipala/frap.git
synced 2024-12-01 00:26:18 +00:00
Proofreading SeparationLogic
This commit is contained in:
parent
7c06dc3541
commit
b7fd72f309
2 changed files with 19 additions and 19 deletions
|
@ -327,11 +327,11 @@ Inductive hoare_triple : forall {result}, hprop -> cmd result -> (result -> hpro
|
||||||
| HtFrame : forall {result} (c : cmd result) P Q R,
|
| HtFrame : forall {result} (c : cmd result) P Q R,
|
||||||
hoare_triple P c Q
|
hoare_triple P c Q
|
||||||
-> hoare_triple (P * R)%sep c (fun r => Q r * R)%sep.
|
-> hoare_triple (P * R)%sep c (fun r => Q r * R)%sep.
|
||||||
(* The *frame rule* is the other big idea of separation. We can extend any
|
(* The *frame rule* is the other big idea of separation logic. We can extend
|
||||||
* Hoare triple by starring an arbitrary assertion [R] into both precondition
|
* any Hoare triple by starring an arbitrary assertion [R] into both
|
||||||
* and postcondition. Note that rule [HtRead] built in a variant of the frame
|
* precondition and postcondition. Note that rule [HtRead] built in a variant
|
||||||
* rule, where the frame predicate [R] may depend on the value [v] being read.
|
* of the frame rule, where the frame predicate [R] may depend on the value [v]
|
||||||
* The other operations can use this generic frame rule instead. *)
|
* being read. The other operations can use this generic frame rule instead. *)
|
||||||
|
|
||||||
|
|
||||||
Notation "{{ P }} c {{ r ~> Q }}" :=
|
Notation "{{ P }} c {{ r ~> Q }}" :=
|
||||||
|
@ -1266,7 +1266,7 @@ Qed.
|
||||||
Fixpoint linkedListSegment (p : nat) (ls : list nat) (q : nat) :=
|
Fixpoint linkedListSegment (p : nat) (ls : list nat) (q : nat) :=
|
||||||
match ls with
|
match ls with
|
||||||
| nil => [| p = q |]
|
| nil => [| p = q |]
|
||||||
| x :: ls' => [| p <> 0 |] * exists p', p |-> x * (p+1) |-> p' * linkedListSegment p' ls' q
|
| x :: ls' => [| p <> 0 |] * exists p', p |--> [x; p'] * linkedListSegment p' ls' q
|
||||||
end%sep.
|
end%sep.
|
||||||
|
|
||||||
(* Next, two [linkedListSegment] lemmas analogous to those for [linkedList]
|
(* Next, two [linkedListSegment] lemmas analogous to those for [linkedList]
|
||||||
|
@ -1280,7 +1280,7 @@ Qed.
|
||||||
|
|
||||||
Lemma linkedListSegment_append : forall q r x ls p,
|
Lemma linkedListSegment_append : forall q r x ls p,
|
||||||
q <> 0
|
q <> 0
|
||||||
-> linkedListSegment p ls q * q |-> x * (q+1) |-> r ===> linkedListSegment p (ls ++ x :: nil) r.
|
-> linkedListSegment p ls q * q |--> [x; r] ===> linkedListSegment p (ls ++ x :: nil) r.
|
||||||
Proof.
|
Proof.
|
||||||
induct ls; cancel; auto.
|
induct ls; cancel; auto.
|
||||||
subst; cancel.
|
subst; cancel.
|
||||||
|
|
Loading…
Reference in a new issue