Using Hoare logic we can derive statements like this, which capture the specification of the function that increments a variable:
{ x = n } inc(x) { x = n + 1}
We'd like to be able to use such statements to construct proofs of statements
that involve procedures like inc()
, defined as follows:
{x = n ∧ y = m } inc(x); inc(y); {x = n+1 ∧ y = m+1}
The problem is that we need an intermediate proposition that is stronger than
what the specification for inc()
gives us directly. One approach is to somehow
say that things other than x
aren't modified by the first increment.
There have been a number of attempts to extend the logic to support modular reasoning about heaps and procedures. Of these, separation logic is probably the most popular.
Separation logic has a similar structure to Hoare logic, but the meaning of {P}c{Q} is different. Semantically, ∀h1, ∀h2, if P(h1), then running c in heap (h1 ⨄ h2) yields a heap h3 = h4 ⨄ h2 where Q(h4). Note that by ⨄ here we mean disjoint union.
Specifications in separation logic generally have the following syntax:
emp the heap is empty x ↦ v the heap contains a single mapping from x to v P * Q * is the _separating conjunction_ of separation logic. Holds on heap h = h1 ⨄ h2 where P(h1) and Q(h2). P ∧ Q Usual semantics P -* Q The "magic wand" implication. Idea: a heap satisfies it if it satisfies Q when extended with some disjoint heap that satisfies P. This is useful for weakest preconditions.
This is nice because the rules for assignment, consequence etc. still hold:
{x ↦ _ } x := v { x ↦ v}
just like in standard Hoare Logic.
Require Import Eqdep.Require Import List. Set Implicit Arguments. Axiom proof_irrelevance : forall (P:Prop) (H1 H2:P), H1 = H2.
In this lecture, we will develop a separation logic for reasoning about imperative Coq programs. Separation logic gives us a crucial principle for modularly reasoning about programs -- the frame rule.
Module FunctionalSepIMP. Definition ptr := nat. Definition ptr_eq := PeanoNat.Nat.eq_dec. Definition le_gt_dec := Arith.Compare_dec.le_gt_dec.
As a demonstration, let's fix the universe of storable types by giving some names to a few standard types as well as an interpretation.
Inductive stype : Set := | Nat_t : stype | Pair_t : stype -> stype -> stype | Sum_t : stype -> stype -> stype | Fn_t : stype -> stype -> stype.decide equality. Defined. Fixpoint interp(t:stype) : Set := (match t with | Nat_t => nat | Pair_t t1 t2 => (interp t1) * (interp t2) | Sum_t t1 t2 => (interp t1) + (interp t2) | Fn_t t1 t2 => (interp t1) -> (interp t2) end)%type.t1, t2: stype{t1 = t2} + {t1 <> t2}
We're going to need to store dynamic values -- a pair of an stype t
and a
value of type interp t
.
Definition dynamic := sigT interp.
We will continue to use lists of pointers and values as the model for heaps. However, to support an easy definition of disjoint union, we will impose the additional constraint that the list is sorted in (strictly) increasing order. It's possible to capture this constraint by defining heaps as a sigma type, where we include a proof that the heap is sorted. That makes some things easier (e.g., arguing that disjoint union is commutative) but makes other things harder. For example, equality on sigmas would demand we need to compare proofs, and use proof-irrelevance. To avoid this we can put in well-formedness constraints in just the right places.
Definition heap := list (ptr * dynamic). Definition empty_heap : heap := nil.
Here is a predicate which says that each pointer in h
is greater than x
.
Fixpoint list_greater (x:ptr) (h:heap) : Prop := match h with | nil => True | (y,v)::h' => x < y /\ list_greater x h' end.
A heap is well-formed if each pointer is less than the rest of the heap, and the rest of the heap is well-formed.
Fixpoint wf (h:heap) : Prop := match h with | nil => True | (x,v)::h' => list_greater x h' /\ wf h' end. Fixpoint indom (x:ptr) (h:heap) : Prop := match h with | nil => False | (y,_)::h' => if ptr_eq x y then True else indom x h' end.
A pointer x
is fresh for h
when it's not in the domain of h
.
Definition fresh x (h:heap) : Prop := ~ indom x h.
Here are standard lookup
and remove
functions on heaps.
Fixpoint lookup (x:ptr) (h:heap) : option dynamic := match h with | nil => None | (y,v)::h' => if ptr_eq x y then Some v else lookup x h' end. Fixpoint remove (x:ptr) (h:heap) : heap := match h with | nil => nil | (y,v)::h' => if ptr_eq x y then h' else (y,v)::(remove x h') end.
Two heaps are disjoint when each pointer in h1
is fresh for h2
.
Fixpoint disjoint (h1 h2:heap) : Prop := match h1 with | nil => True | (x,_)::h1' => ~indom x h2 /\ disjoint h1' h2 end.
To insert
into a heap, we use insertion sort.
Fixpoint insert (x:nat) (v:dynamic) (h:heap) : heap := match h with | nil => (x,v)::nil | (y,w)::h' => if le_gt_dec x y then (x,v)::(y,w)::h' else (y,w)::(insert x v h') end.
We can merge two heaps using our insert
function.
Definition merge (h1 h2:heap) : heap := List.fold_right (fun p h => insert (fst p) (snd p) h) h2 h1.
As in earlier lectures, we model commands via shallow embedding and using a monad. That is, a command takes a heap and returns an optional heap and a result.
Definition Cmd(t:Type) := heap -> option(heap * t). Definition ret t (x:t) : Cmd t := fun h => Some (h,x). Definition exit t : Cmd t := fun h => None. Definition bind t u (c:Cmd t) (f:t -> Cmd u) : Cmd u := fun h1 => match c h1 with | None => None | Some (h2,v) => f v h2 end. Declare Scope cmd_scope.
We introduce some notation similar to Haskell's "do" notation.
Notation "x <- c ; f" := (bind c (fun x => f)) (right associativity, at level 84) : cmd_scope. Notation "c ;; f" := (bind c (fun _:unit => f)) (right associativity, at level 84) : cmd_scope. Local Open Scope cmd_scope. Definition run(t:Type)(c:Cmd t) := c empty_heap.
The untyped read operation returns a dynamic value.
Definition untyped_read (p:ptr) : Cmd dynamic := fun h => match lookup p h with | None => None | Some u => Some (h,u) end.
This abbreviation may look funny (since we don't use t
), but it will make it
easier for Coq to figure out the stype
at which we're using things. Note
that this is a little misleading because in general, a ptr
can map to a
value of any stype!
Definition tptr (t:stype) := ptr.
The typed read operation first does an untyped read, and then checks that the
stype
matches what we expected, failing if not, and returning the underlying
value otherwise.
t: stype
p: tptr tCmd (interp t)t0: stype
p0: tptr t0
t: stype
p: ptr
d: dynamic
t': stype
v: interp t'
Heq: t = t'interp tapply v. Defined.t0: stype
p0: tptr t0
p: ptr
d: dynamic
t': stype
v: interp t'interp t'
Likewise, we have untyped and typed write operations.
Definition untyped_write (p:ptr) (v:dynamic) : Cmd unit := fun h => match lookup p h with | None => None | Some _ => Some (insert p v (remove p h), tt) end. Definition write (t:stype) (p:tptr t) (v:interp t) : Cmd unit := untyped_write p (existT interp t v).
To allocate a new pointer, we need to pick something fresh for the heap. So we choose the maximum value in the heap plus one.
Definition max (p1 p2:ptr) := if le_gt_dec p1 p2 then p2 else p1. Definition max_heap(h:heap) := List.fold_right (fun p n => max (fst p) n) 0 h. Definition untyped_new (v:dynamic) : Cmd ptr := fun h => let p := 1 + max_heap h in Some (insert p v h, p). Definition new(t:stype)(v:interp t) : Cmd (tptr t) := untyped_new (existT interp t v). Definition free (p:ptr) : Cmd unit := fun h => match lookup p h with | None => None | Some _ => Some(remove p h, tt) end.
An hprop
is a predicate on heaps.
Definition hprop := heap -> Prop.
emp
holds only when the heap is empty. One way to think of emp
is that as
a pre-condition, it tells us that we don't have the right to access anything in
the heap. emp
plays the role of a unit for star
which is defined below.
Definition emp : hprop := fun h => h = nil.
top
holds on any heap.
Definition top : hprop := fun _ => True. Declare Scope sep_scope.
x ---> d
holds when the heap contains exactly one pointer x
which
points to d
. So we can think of this predicate as permission to
read or write x
and if we read it, we'll get out the dynamic value d
.
However, we're not allowed to read or write any other location if this
is our only pre-condition.
Definition ptsto (x:ptr) (d:dynamic) : hprop := fun h => h = (x,d)::nil. Infix "--->" := ptsto (at level 70) : sep_scope. Local Open Scope sep_scope.
x --> v
is the same as above, except we make the type explicit.
Definition typed_ptsto (t:stype) (x:tptr t) (v:interp t) : hprop := x ---> (existT interp t v). Arguments typed_ptsto {t}. Infix "-->" := typed_ptsto (at level 70) : sep_scope.
P ** Q
holds when the heap can be broken into disjoint heaps h1
and h2
such that h1
satisfies P
and h2
satisfies Q
. For example x --> v ** y --> v
tells us that we have the right to access both x
and y
, but that x <> y
.
Definition star (P Q:hprop) : hprop := fun h => exists h1:heap, exists h2:heap, wf h1 /\ wf h2 /\ P h1 /\ Q h2 /\ h = merge h1 h2 /\ disjoint h1 h2. Infix "**" := star (right associativity, at level 80) : sep_scope.
When P
is a pure predicate in the sense that it doesn't refer to a
heap, then we can use pure P
as a predicate on the heap. Note that
we require that the heap that it corresponds to is empty.
Definition pure (P:Prop) : hprop := fun h => emp h /\ P.
sing h
is the singleton predciate on heaps -- i.e., it holds on h'
only when h'
is equal to h
.
Definition sing (h:heap) : hprop := fun h' => h = h'.
This definition lifts an existential quantifier up to a predicate on heaps.
Definition hexists (T:Type) (f:T -> hprop) : hprop := fun h => exists x:T, f x h.
So for instance, we can define x -->?
as follows
Local Open Scope sep_scope. Definition ptsto_some (x:ptr) := hexists (fun v => x ---> v). Notation "x '-->?'" := (ptsto_some x) (at level 70) : sep_scope.
This notation will be useful for reasoning about when one separation predicate can be used in lieu of another -- i.e., a form of implication over heaps. Note that this is not the "magic-wand" of separation logic, but rather, a meta-level implication.
Definition himp(P Q:hprop) : Prop := forall h, wf h -> P h -> Q h. Infix "==>" := himp (at level 90) : sep_scope.
Here is a standard Greg (Morrisett) simplification tactic.
Ltac mysimp := repeat (simpl ; match goal with | [ p : (_ * _)%type |- _ ] => destruct p | [ H : _ /\ _ |- _ ] => destruct H | [ H : exists _, _ |- _ ] => destruct H | [ |- context[ptr_eq ?x ?y] ] => destruct (ptr_eq x y) ; try subst ; try congruence | [ H : context[ptr_eq ?x ?y] |- _ ] => destruct (ptr_eq x y) ; try subst ; try congruence | [ |- context[le_gt_dec ?x ?y] ] => destruct (le_gt_dec x y) ; try lia | [ H : context[le_gt_dec ?x ?y] |- _ ] => destruct (le_gt_dec x y) ; try lia | [ |- _ /\ _ ] => split | [ H : (_,_) = (_,_) |- _ ] => injection H ; clear H ; intros ; try subst | [ H : Some _ = Some _ |- _ ] => injection H ; clear H ; intros ; try subst | [ H : Some _ = None |- _ ] => congruence | [ H : None = Some _ |- _ ] => congruence | [ H : ?x <> ?x |- _ ] => congruence | [ H : ~ True |- _ ] => contradiction H ; auto | [ H : False |- _ ] => contradiction H ; auto | [ |- forall _, _ ] => intros | _ => auto end).
We need a ton of lemmas for reasoning about heaps. The keys ones are showing that various operations preserve well-formedness, or that merging well-formed heaps is commutative and associative, etc.
For example, removing a pointer keeps a heap well-formed.
x: ptr
h: heapwf h -> wf (remove x h)x: ptr
h: heapwf h -> wf (remove x h)generalize h H ; induction h0 ; mysimp. Qed.x, p: ptr
d: dynamic
h: list (ptr * dynamic)
IHh: wf h -> wf (remove x h)
n: x <> p
H: list_greater p h
H0: wf hlist_greater p (remove x h)forall (h1 h2 : heap) (p : ptr), disjoint h1 h2 -> indom p h1 -> ~ indom p h2forall (h1 h2 : heap) (p : ptr), disjoint h1 h2 -> indom p h1 -> ~ indom p h2forall (h2 : heap) (p : ptr), disjoint nil h2 -> indom p nil -> ~ indom p h2a: (ptr * dynamic)%type
h1: list (ptr * dynamic)
IHh1: forall (h2 : heap) (p : ptr), disjoint h1 h2 -> indom p h1 -> ~ indom p h2forall (h2 : heap) (p : ptr), disjoint (a :: h1) h2 -> indom p (a :: h1) -> ~ indom p h2mysimp.forall (h2 : heap) (p : ptr), disjoint nil h2 -> indom p nil -> ~ indom p h2a: (ptr * dynamic)%type
h1: list (ptr * dynamic)
IHh1: forall (h2 : heap) (p : ptr), disjoint h1 h2 -> indom p h1 -> ~ indom p h2forall (h2 : heap) (p : ptr), disjoint (a :: h1) h2 -> indom p (a :: h1) -> ~ indom p h2a: (ptr * dynamic)%type
h1: list (ptr * dynamic)
IHh1: forall (h2 : heap) (p : ptr), disjoint h1 h2 -> indom p h1 -> ~ indom p h2
h2: heap
p: ptr
H: disjoint (a :: h1) h2
H0: indom p (a :: h1)~ indom p h2p1: ptr
d1: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : heap) (p : ptr), disjoint h1 h2 -> indom p h1 -> ~ indom p h2
h2: heap
p: ptr
H: disjoint ((p1, d1) :: h1) h2
H0: indom p ((p1, d1) :: h1)~ indom p h2p1: ptr
d1: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : heap) (p : ptr), disjoint h1 h2 -> indom p h1 -> ~ indom p h2
h2: heap
p: ptr
H: disjoint ((p1, d1) :: h1) h2
H0: indom p ((p1, d1) :: h1)
e: p = p1~ indom p h2p1: ptr
d1: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : heap) (p : ptr), disjoint h1 h2 -> indom p h1 -> ~ indom p h2
h2: heap
p: ptr
H: disjoint ((p1, d1) :: h1) h2
H0: indom p ((p1, d1) :: h1)
n: p <> p1~ indom p h2p1: ptr
d1: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : heap) (p : ptr), disjoint h1 h2 -> indom p h1 -> ~ indom p h2
h2: heap
p: ptr
H: disjoint ((p1, d1) :: h1) h2
H0: indom p ((p1, d1) :: h1)
e: p = p1~ indom p h2p1: ptr
d1: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : heap) (p : ptr), disjoint h1 h2 -> indom p h1 -> ~ indom p h2
h2: heap
p: ptr
H: disjoint ((p, d1) :: h1) h2
H0: indom p ((p1, d1) :: h1)
e: p = p1~ indom p h2mysimp.p1: ptr
d1: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : heap) (p : ptr), disjoint h1 h2 -> indom p h1 -> ~ indom p h2
h2: heap
p: ptr
H: ~ indom p h2 /\ (fix disjoint (h1 h2 : heap) {struct h1} : Prop := match h1 with | nil => True | (x, _) :: h1' => ~ indom x h2 /\ disjoint h1' h2 end) h1 h2
H0: indom p ((p1, d1) :: h1)
e: p = p1~ indom p h2p1: ptr
d1: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : heap) (p : ptr), disjoint h1 h2 -> indom p h1 -> ~ indom p h2
h2: heap
p: ptr
H: disjoint ((p1, d1) :: h1) h2
H0: indom p ((p1, d1) :: h1)
n: p <> p1~ indom p h2p1: ptr
d1: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : heap) (p : ptr), disjoint h1 h2 -> indom p h1 -> ~ indom p h2
h2: heap
p: ptr
H: disjoint ((p1, d1) :: h1) h2
H0: indom p ((p1, d1) :: h1)
n: p <> p1disjoint h1 h2p1: ptr
d1: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : heap) (p : ptr), disjoint h1 h2 -> indom p h1 -> ~ indom p h2
h2: heap
p: ptr
H: disjoint ((p1, d1) :: h1) h2
H0: indom p ((p1, d1) :: h1)
n: p <> p1indom p h1p1: ptr
d1: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : heap) (p : ptr), disjoint h1 h2 -> indom p h1 -> ~ indom p h2
h2: heap
p: ptr
H: ~ indom p1 h2 /\ (fix disjoint (h1 h2 : heap) {struct h1} : Prop := match h1 with | nil => True | (x, _) :: h1' => ~ indom x h2 /\ disjoint h1' h2 end) h1 h2
H0: indom p ((p1, d1) :: h1)
n: p <> p1disjoint h1 h2p1: ptr
d1: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : heap) (p : ptr), disjoint h1 h2 -> indom p h1 -> ~ indom p h2
h2: heap
p: ptr
H: disjoint ((p1, d1) :: h1) h2
H0: indom p ((p1, d1) :: h1)
n: p <> p1indom p h1p1: ptr
d1: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : heap) (p : ptr), disjoint h1 h2 -> indom p h1 -> ~ indom p h2
h2: heap
p: ptr
H: disjoint ((p1, d1) :: h1) h2
H0: indom p ((p1, d1) :: h1)
n: p <> p1indom p h1mysimp. Qed.p1: ptr
d1: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : heap) (p : ptr), disjoint h1 h2 -> indom p h1 -> ~ indom p h2
h2: heap
p: ptr
H: disjoint ((p1, d1) :: h1) h2
H0: if ptr_eq p p1 then True else (fix indom (x : ptr) (h : heap) {struct h} : Prop := match h with | nil => False | (y, _) :: h' => if ptr_eq x y then True else indom x h' end) p h1
n: p <> p1indom p h1forall (h : heap) (p0 p : ptr) (d : dynamic), ~ indom p0 ((p, d) :: h) -> ~ indom p0 hmysimp. Qed.forall (h : heap) (p0 p : ptr) (d : dynamic), ~ indom p0 ((p, d) :: h) -> ~ indom p0 hforall (h1 : heap) (h2 : list (ptr * dynamic)) (p : ptr * dynamic), disjoint h1 (p :: h2) -> disjoint h1 h2forall (h1 : heap) (h2 : list (ptr * dynamic)) (p : ptr * dynamic), disjoint h1 (p :: h2) -> disjoint h1 h2forall (h2 : list (ptr * dynamic)) (p : ptr * dynamic), disjoint nil (p :: h2) -> disjoint nil h2a: (ptr * dynamic)%type
h1: list (ptr * dynamic)
IHh1: forall (h2 : list (ptr * dynamic)) (p : ptr * dynamic), disjoint h1 (p :: h2) -> disjoint h1 h2forall (h2 : list (ptr * dynamic)) (p : ptr * dynamic), disjoint (a :: h1) (p :: h2) -> disjoint (a :: h1) h2forall (h2 : list (ptr * dynamic)) (p : ptr * dynamic), disjoint nil (p :: h2) -> disjoint nil h2mysimp.h2: list (ptr * dynamic)
p: (ptr * dynamic)%type
H: disjoint nil (p :: h2)disjoint nil h2a: (ptr * dynamic)%type
h1: list (ptr * dynamic)
IHh1: forall (h2 : list (ptr * dynamic)) (p : ptr * dynamic), disjoint h1 (p :: h2) -> disjoint h1 h2forall (h2 : list (ptr * dynamic)) (p : ptr * dynamic), disjoint (a :: h1) (p :: h2) -> disjoint (a :: h1) h2a: (ptr * dynamic)%type
h1: list (ptr * dynamic)
IHh1: forall (h2 : list (ptr * dynamic)) (p : ptr * dynamic), disjoint h1 (p :: h2) -> disjoint h1 h2
h2: list (ptr * dynamic)
p: (ptr * dynamic)%type
H: disjoint (a :: h1) (p :: h2)disjoint (a :: h1) h2p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : list (ptr * dynamic)) (p : ptr * dynamic), disjoint h1 (p :: h2) -> disjoint h1 h2
h2: list (ptr * dynamic)
p: ptr
d: dynamic
H: disjoint ((p0, d0) :: h1) ((p, d) :: h2)~ indom p0 h2p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : list (ptr * dynamic)) (p : ptr * dynamic), disjoint h1 (p :: h2) -> disjoint h1 h2
h2: list (ptr * dynamic)
p: ptr
d: dynamic
H: disjoint ((p0, d0) :: h1) ((p, d) :: h2)disjoint h1 h2p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : list (ptr * dynamic)) (p : ptr * dynamic), disjoint h1 (p :: h2) -> disjoint h1 h2
h2: list (ptr * dynamic)
p: ptr
d: dynamic
H: disjoint ((p0, d0) :: h1) ((p, d) :: h2)~ indom p0 ((p, d) :: h2)p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : list (ptr * dynamic)) (p : ptr * dynamic), disjoint h1 (p :: h2) -> disjoint h1 h2
h2: list (ptr * dynamic)
p: ptr
d: dynamic
H: disjoint ((p0, d0) :: h1) ((p, d) :: h2)disjoint h1 h2p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : list (ptr * dynamic)) (p : ptr * dynamic), disjoint h1 (p :: h2) -> disjoint h1 h2
h2: list (ptr * dynamic)
p: ptr
d: dynamic
H: ~ indom p0 ((p, d) :: h2) /\ (fix disjoint (h1 h2 : heap) {struct h1} : Prop := match h1 with | nil => True | (x, _) :: h1' => ~ indom x h2 /\ disjoint h1' h2 end) h1 ((p, d) :: h2)~ indom p0 ((p, d) :: h2)p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : list (ptr * dynamic)) (p : ptr * dynamic), disjoint h1 (p :: h2) -> disjoint h1 h2
h2: list (ptr * dynamic)
p: ptr
d: dynamic
H: disjoint ((p0, d0) :: h1) ((p, d) :: h2)disjoint h1 h2p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : list (ptr * dynamic)) (p : ptr * dynamic), disjoint h1 (p :: h2) -> disjoint h1 h2
h2: list (ptr * dynamic)
p: ptr
d: dynamic
H: ~ indom p0 ((p, d) :: h2)
H0: (fix disjoint (h1 h2 : heap) {struct h1} : Prop := match h1 with | nil => True | (x, _) :: h1' => ~ indom x h2 /\ disjoint h1' h2 end) h1 ((p, d) :: h2)~ indom p0 ((p, d) :: h2)p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : list (ptr * dynamic)) (p : ptr * dynamic), disjoint h1 (p :: h2) -> disjoint h1 h2
h2: list (ptr * dynamic)
p: ptr
d: dynamic
H: disjoint ((p0, d0) :: h1) ((p, d) :: h2)disjoint h1 h2p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : list (ptr * dynamic)) (p : ptr * dynamic), disjoint h1 (p :: h2) -> disjoint h1 h2
h2: list (ptr * dynamic)
p: ptr
d: dynamic
H: disjoint ((p0, d0) :: h1) ((p, d) :: h2)disjoint h1 h2p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : list (ptr * dynamic)) (p : ptr * dynamic), disjoint h1 (p :: h2) -> disjoint h1 h2
h2: list (ptr * dynamic)
p: ptr
d: dynamic
H: disjoint ((p0, d0) :: h1) ((p, d) :: h2)disjoint h1 ((p, d) :: h2)mysimp. Qed.p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall (h2 : list (ptr * dynamic)) (p : ptr * dynamic), disjoint h1 (p :: h2) -> disjoint h1 h2
h2: list (ptr * dynamic)
p: ptr
d: dynamic
H: ~ indom p0 ((p, d) :: h2) /\ (fix disjoint (h1 h2 : heap) {struct h1} : Prop := match h1 with | nil => True | (x, _) :: h1' => ~ indom x h2 /\ disjoint h1' h2 end) h1 ((p, d) :: h2)disjoint h1 ((p, d) :: h2)
Disjointness is commutative.
forall h1 h2 : heap, disjoint h1 h2 -> disjoint h2 h1forall h1 h2 : heap, disjoint h1 h2 -> disjoint h2 h1p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 : heap, disjoint h1 h2 -> disjoint h2 h1
p: ptr
d: dynamic
h2: list (ptr * dynamic)
IHh2: disjoint ((p0, d0) :: h1) h2 -> disjoint h2 ((p0, d0) :: h1)
n: p0 <> p
n0: p <> p0
H: ~ indom p0 h2
H0: disjoint h1 ((p, d) :: h2)~ indom p h1p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 : heap, disjoint h1 h2 -> disjoint h2 h1
p: ptr
d: dynamic
h2: list (ptr * dynamic)
IHh2: disjoint ((p0, d0) :: h1) h2 -> disjoint h2 ((p0, d0) :: h1)
n: p0 <> p
n0: p <> p0
H: ~ indom p0 h2
H0: disjoint h1 ((p, d) :: h2)disjoint h2 ((p0, d0) :: h1)p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 : heap, disjoint h1 h2 -> disjoint h2 h1
p: ptr
d: dynamic
h2: list (ptr * dynamic)
IHh2: disjoint ((p0, d0) :: h1) h2 -> disjoint h2 ((p0, d0) :: h1)
n: p0 <> p
n0: p <> p0
H: ~ indom p0 h2
H0: disjoint h1 ((p, d) :: h2)disjoint ((p, d) :: h2) h1p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 : heap, disjoint h1 h2 -> disjoint h2 h1
p: ptr
d: dynamic
h2: list (ptr * dynamic)
IHh2: disjoint ((p0, d0) :: h1) h2 -> disjoint h2 ((p0, d0) :: h1)
n: p0 <> p
n0: p <> p0
H: ~ indom p0 h2
H0: disjoint h1 ((p, d) :: h2)indom p ((p, d) :: h2)p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 : heap, disjoint h1 h2 -> disjoint h2 h1
p: ptr
d: dynamic
h2: list (ptr * dynamic)
IHh2: disjoint ((p0, d0) :: h1) h2 -> disjoint h2 ((p0, d0) :: h1)
n: p0 <> p
n0: p <> p0
H: ~ indom p0 h2
H0: disjoint h1 ((p, d) :: h2)disjoint h2 ((p0, d0) :: h1)p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 : heap, disjoint h1 h2 -> disjoint h2 h1
p: ptr
d: dynamic
h2: list (ptr * dynamic)
IHh2: disjoint ((p0, d0) :: h1) h2 -> disjoint h2 ((p0, d0) :: h1)
n: p0 <> p
n0: p <> p0
H: ~ indom p0 h2
H0: disjoint h1 ((p, d) :: h2)indom p ((p, d) :: h2)p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 : heap, disjoint h1 h2 -> disjoint h2 h1
p: ptr
d: dynamic
h2: list (ptr * dynamic)
IHh2: disjoint ((p0, d0) :: h1) h2 -> disjoint h2 ((p0, d0) :: h1)
n: p0 <> p
n0: p <> p0
H: ~ indom p0 h2
H0: disjoint h1 ((p, d) :: h2)disjoint h2 ((p0, d0) :: h1)p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 : heap, disjoint h1 h2 -> disjoint h2 h1
p: ptr
d: dynamic
h2: list (ptr * dynamic)
IHh2: disjoint ((p0, d0) :: h1) h2 -> disjoint h2 ((p0, d0) :: h1)
n: p0 <> p
n0: p <> p0
H: ~ indom p0 h2
H0: disjoint h1 ((p, d) :: h2)disjoint h2 ((p0, d0) :: h1)p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 : heap, disjoint h1 h2 -> disjoint h2 h1
p: ptr
d: dynamic
h2: list (ptr * dynamic)
IHh2: disjoint ((p0, d0) :: h1) h2 -> disjoint h2 ((p0, d0) :: h1)
n: p0 <> p
n0: p <> p0
H: ~ indom p0 h2
H0: disjoint h1 ((p, d) :: h2)disjoint ((p0, d0) :: h1) h2p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 : heap, disjoint h1 h2 -> disjoint h2 h1
p: ptr
d: dynamic
h2: list (ptr * dynamic)
IHh2: disjoint ((p0, d0) :: h1) h2 -> disjoint h2 ((p0, d0) :: h1)
n: p0 <> p
n0: p <> p0
H: ~ indom p0 h2
H0: disjoint h1 ((p, d) :: h2)disjoint h1 h2mysimp. Qed.p0: ptr
d0: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 : heap, disjoint h1 h2 -> disjoint h2 h1
p: ptr
d: dynamic
h2: list (ptr * dynamic)
IHh2: disjoint ((p0, d0) :: h1) h2 -> disjoint h2 ((p0, d0) :: h1)
n: p0 <> p
n0: p <> p0
H: ~ indom p0 h2
H0: disjoint h1 ((p, d) :: h2)disjoint h1 ((p, d) :: h2)x, y: natx < y -> forall h : heap, list_greater y h -> list_greater x hlia. Qed.x, y: nat
H: x < y
p: ptr
d: dynamic
h: list (ptr * dynamic)
IHh: list_greater y h -> list_greater x h
H0: y < p
H1: list_greater y hx < p
Inserting a fresh pointer keeps a heap well-formed.
h: heapwf h -> forall (x : ptr) (v : dynamic), fresh x h -> wf (insert x v h)h: heapwf h -> forall (x : ptr) (v : dynamic), fresh x h -> wf (insert x v h)p: ptr
d: dynamic
h: list (ptr * dynamic)
IHh: wf h -> forall (x : ptr) (v : dynamic), ~ indom x h -> wf (insert x v h)
H: list_greater p h
H1: wf h
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x h
l: x <= px < pp: ptr
d: dynamic
h: list (ptr * dynamic)
IHh: wf h -> forall (x : ptr) (v : dynamic), ~ indom x h -> wf (insert x v h)
H: list_greater p h
H1: wf h
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x h
l: x <= plist_greater x hp: ptr
d: dynamic
h: list (ptr * dynamic)
IHh: wf h -> forall (x : ptr) (v : dynamic), ~ indom x h -> wf (insert x v h)
H: list_greater p h
H1: wf h
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x h
g: x > plist_greater p (insert x v h)p: ptr
d: dynamic
h: list (ptr * dynamic)
IHh: wf h -> forall (x : ptr) (v : dynamic), ~ indom x h -> wf (insert x v h)
H: list_greater p h
H1: wf h
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x h
l: x <= plist_greater x hp: ptr
d: dynamic
h: list (ptr * dynamic)
IHh: wf h -> forall (x : ptr) (v : dynamic), ~ indom x h -> wf (insert x v h)
H: list_greater p h
H1: wf h
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x h
g: x > plist_greater p (insert x v h)p: ptr
d: dynamic
h: list (ptr * dynamic)
IHh: wf h -> forall (x : ptr) (v : dynamic), ~ indom x h -> wf (insert x v h)
H: list_greater p h
H1: wf h
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x h
l: x <= px < pp: ptr
d: dynamic
h: list (ptr * dynamic)
IHh: wf h -> forall (x : ptr) (v : dynamic), ~ indom x h -> wf (insert x v h)
H: list_greater p h
H1: wf h
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x h
g: x > plist_greater p (insert x v h)p: ptr
d: dynamic
h: list (ptr * dynamic)
IHh: wf h -> forall (x : ptr) (v : dynamic), ~ indom x h -> wf (insert x v h)
H: list_greater p h
H1: wf h
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x h
g: x > plist_greater p (insert x v h)p: ptr
d: dynamic
h: list (ptr * dynamic)
IHh: wf h -> forall (x : ptr) (v : dynamic), ~ indom x h -> wf (insert x v h)
H: list_greater p h
H1: wf h
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x h
g: x > pwf (insert x v h) -> list_greater p (insert x v h)p: ptr
d: dynamic
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: wf ((p0, d0) :: h) -> forall (x : ptr) (v : dynamic), ~ indom x ((p0, d0) :: h) -> wf (insert x v ((p0, d0) :: h))
H: list_greater p ((p0, d0) :: h)
H1: wf ((p0, d0) :: h)
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x ((p0, d0) :: h)
g: x > p
l: x <= p0
H2: x < p0
H5: list_greater x h
H3: list_greater p0 h
H4: wf hp < p0p: ptr
d: dynamic
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: wf ((p0, d0) :: h) -> forall (x : ptr) (v : dynamic), ~ indom x ((p0, d0) :: h) -> wf (insert x v ((p0, d0) :: h))
H: list_greater p ((p0, d0) :: h)
H1: wf ((p0, d0) :: h)
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x ((p0, d0) :: h)
g: x > p
l: x <= p0
H2: x < p0
H5: list_greater x h
H3: list_greater p0 h
H4: wf hlist_greater p hp: ptr
d: dynamic
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: wf ((p0, d0) :: h) -> forall (x : ptr) (v : dynamic), ~ indom x ((p0, d0) :: h) -> wf (insert x v ((p0, d0) :: h))
H: list_greater p ((p0, d0) :: h)
H1: wf ((p0, d0) :: h)
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x ((p0, d0) :: h)
g: x > p
g0: x > p0
H2: list_greater p0 (insert x v h)
H3: wf (insert x v h)p < p0p: ptr
d: dynamic
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: wf ((p0, d0) :: h) -> forall (x : ptr) (v : dynamic), ~ indom x ((p0, d0) :: h) -> wf (insert x v ((p0, d0) :: h))
H: list_greater p ((p0, d0) :: h)
H1: wf ((p0, d0) :: h)
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x ((p0, d0) :: h)
g: x > p
g0: x > p0
H2: list_greater p0 (insert x v h)
H3: wf (insert x v h)list_greater p (insert x v h)p: ptr
d: dynamic
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: wf ((p0, d0) :: h) -> forall (x : ptr) (v : dynamic), ~ indom x ((p0, d0) :: h) -> wf (insert x v ((p0, d0) :: h))
H: list_greater p ((p0, d0) :: h)
H1: wf ((p0, d0) :: h)
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x ((p0, d0) :: h)
g: x > p
l: x <= p0
H2: x < p0
H5: list_greater x h
H3: list_greater p0 h
H4: wf hlist_greater p hp: ptr
d: dynamic
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: wf ((p0, d0) :: h) -> forall (x : ptr) (v : dynamic), ~ indom x ((p0, d0) :: h) -> wf (insert x v ((p0, d0) :: h))
H: list_greater p ((p0, d0) :: h)
H1: wf ((p0, d0) :: h)
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x ((p0, d0) :: h)
g: x > p
g0: x > p0
H2: list_greater p0 (insert x v h)
H3: wf (insert x v h)p < p0p: ptr
d: dynamic
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: wf ((p0, d0) :: h) -> forall (x : ptr) (v : dynamic), ~ indom x ((p0, d0) :: h) -> wf (insert x v ((p0, d0) :: h))
H: list_greater p ((p0, d0) :: h)
H1: wf ((p0, d0) :: h)
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x ((p0, d0) :: h)
g: x > p
g0: x > p0
H2: list_greater p0 (insert x v h)
H3: wf (insert x v h)list_greater p (insert x v h)p: ptr
d: dynamic
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: wf ((p0, d0) :: h) -> forall (x : ptr) (v : dynamic), ~ indom x ((p0, d0) :: h) -> wf (insert x v ((p0, d0) :: h))
H: list_greater p ((p0, d0) :: h)
H1: wf ((p0, d0) :: h)
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x ((p0, d0) :: h)
g: x > p
l: x <= p0
H2: x < p0
H5: list_greater x h
H3: list_greater p0 h
H4: wf hp < p0p: ptr
d: dynamic
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: wf ((p0, d0) :: h) -> forall (x : ptr) (v : dynamic), ~ indom x ((p0, d0) :: h) -> wf (insert x v ((p0, d0) :: h))
H: list_greater p ((p0, d0) :: h)
H1: wf ((p0, d0) :: h)
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x ((p0, d0) :: h)
g: x > p
g0: x > p0
H2: list_greater p0 (insert x v h)
H3: wf (insert x v h)p < p0p: ptr
d: dynamic
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: wf ((p0, d0) :: h) -> forall (x : ptr) (v : dynamic), ~ indom x ((p0, d0) :: h) -> wf (insert x v ((p0, d0) :: h))
H: list_greater p ((p0, d0) :: h)
H1: wf ((p0, d0) :: h)
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x ((p0, d0) :: h)
g: x > p
g0: x > p0
H2: list_greater p0 (insert x v h)
H3: wf (insert x v h)list_greater p (insert x v h)p: ptr
d: dynamic
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: wf ((p0, d0) :: h) -> forall (x : ptr) (v : dynamic), ~ indom x ((p0, d0) :: h) -> wf (insert x v ((p0, d0) :: h))
H: list_greater p ((p0, d0) :: h)
H1: wf ((p0, d0) :: h)
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x ((p0, d0) :: h)
g: x > p
g0: x > p0
H2: list_greater p0 (insert x v h)
H3: wf (insert x v h)p < p0p: ptr
d: dynamic
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: wf ((p0, d0) :: h) -> forall (x : ptr) (v : dynamic), ~ indom x ((p0, d0) :: h) -> wf (insert x v ((p0, d0) :: h))
H: list_greater p ((p0, d0) :: h)
H1: wf ((p0, d0) :: h)
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x ((p0, d0) :: h)
g: x > p
g0: x > p0
H2: list_greater p0 (insert x v h)
H3: wf (insert x v h)list_greater p (insert x v h)p: ptr
d: dynamic
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: wf ((p0, d0) :: h) -> forall (x : ptr) (v : dynamic), ~ indom x ((p0, d0) :: h) -> wf (insert x v ((p0, d0) :: h))
H: p < p0 /\ list_greater p h
H1: wf ((p0, d0) :: h)
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x ((p0, d0) :: h)
g: x > p
g0: x > p0
H2: list_greater p0 (insert x v h)
H3: wf (insert x v h)p < p0p: ptr
d: dynamic
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: wf ((p0, d0) :: h) -> forall (x : ptr) (v : dynamic), ~ indom x ((p0, d0) :: h) -> wf (insert x v ((p0, d0) :: h))
H: list_greater p ((p0, d0) :: h)
H1: wf ((p0, d0) :: h)
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x ((p0, d0) :: h)
g: x > p
g0: x > p0
H2: list_greater p0 (insert x v h)
H3: wf (insert x v h)list_greater p (insert x v h)p: ptr
d: dynamic
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: wf ((p0, d0) :: h) -> forall (x : ptr) (v : dynamic), ~ indom x ((p0, d0) :: h) -> wf (insert x v ((p0, d0) :: h))
H: list_greater p ((p0, d0) :: h)
H1: wf ((p0, d0) :: h)
x: ptr
v: dynamic
n: x <> p
H0: ~ indom x ((p0, d0) :: h)
g: x > p
g0: x > p0
H2: list_greater p0 (insert x v h)
H3: wf (insert x v h)list_greater p (insert x v h)p: ptr
d: dynamic
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: list_greater p0 h /\ wf h -> forall (x : ptr) (v : dynamic), ~ (if ptr_eq x p0 then True else indom x h) -> wf (if le_gt_dec x p0 then (x, v) :: (p0, d0) :: h else (p0, d0) :: insert x v h)
H: p < p0 /\ list_greater p h
H1: list_greater p0 h /\ wf h
x: ptr
v: dynamic
n: x <> p
H0: ~ (if ptr_eq x p0 then True else indom x h)
g: x > p
g0: x > p0
H2: list_greater p0 (insert x v h)
H3: wf (insert x v h)list_greater p (insert x v h)p: ptr
d: dynamic
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: list_greater p0 h /\ wf h -> forall (x : ptr) (v : dynamic), ~ (if ptr_eq x p0 then True else indom x h) -> wf (if le_gt_dec x p0 then (x, v) :: (p0, d0) :: h else (p0, d0) :: insert x v h)
H: p < p0 /\ list_greater p h
H1: list_greater p0 h /\ wf h
x: ptr
v: dynamic
n: x <> p
H0: ~ (if ptr_eq x p0 then True else indom x h)
g: x > p
g0: x > p0
H2: list_greater p0 (insert x v h)
H3: wf (insert x v h)p < p0p: ptr
d: dynamic
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: list_greater p0 h /\ wf h -> forall (x : ptr) (v : dynamic), ~ (if ptr_eq x p0 then True else indom x h) -> wf (if le_gt_dec x p0 then (x, v) :: (p0, d0) :: h else (p0, d0) :: insert x v h)
H: p < p0 /\ list_greater p h
H1: list_greater p0 h /\ wf h
x: ptr
v: dynamic
n: x <> p
H0: ~ (if ptr_eq x p0 then True else indom x h)
g: x > p
g0: x > p0
H2: list_greater p0 (insert x v h)
H3: wf (insert x v h)
H4: p < p0list_greater p (insert x v h)eapply lg_trans ; eauto. Qed.p: ptr
d: dynamic
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: list_greater p0 h /\ wf h -> forall (x : ptr) (v : dynamic), ~ (if ptr_eq x p0 then True else indom x h) -> wf (if le_gt_dec x p0 then (x, v) :: (p0, d0) :: h else (p0, d0) :: insert x v h)
H: p < p0 /\ list_greater p h
H1: list_greater p0 h /\ wf h
x: ptr
v: dynamic
n: x <> p
H0: ~ (if ptr_eq x p0 then True else indom x h)
g: x > p
g0: x > p0
H2: list_greater p0 (insert x v h)
H3: wf (insert x v h)
H4: p < p0list_greater p (insert x v h)p: ptr
h: heaplist_greater p h -> fresh p hp: ptr
h: heaplist_greater p h -> fresh p hlia. Qed.p0: ptr
d: dynamic
h: list (ptr * dynamic)
IHh: list_greater p0 h -> ~ indom p0 h
H: p0 < p0
H0: list_greater p0 h~ Truex: ptr
h: heapfresh x h -> forall (y : ptr) (v : dynamic), y <> x -> fresh x (insert y v h)unfold fresh ; induction h ; mysimp. Qed.x: ptr
h: heapfresh x h -> forall (y : ptr) (v : dynamic), y <> x -> fresh x (insert y v h)h1, h2: heap
p: ptrfresh p h1 -> fresh p h2 -> fresh p (merge h1 h2)h1, h2: heap
p: ptrfresh p h1 -> fresh p h2 -> fresh p (merge h1 h2)apply not_indom_insert ; unfold fresh ; auto. Qed.p0: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p h1 -> ~ indom p h2 -> ~ indom p (merge h1 h2)
n: p <> p0
H: ~ indom p h1
H0: ~ indom p h2~ indom p (insert p0 d (merge h1 h2))
Merging two well-formed, disjoint heaps results in a well-formed heap.
h1, h2: heapwf h1 -> wf h2 -> disjoint h1 h2 -> wf (merge h1 h2)h1, h2: heapwf h1 -> wf h2 -> disjoint h1 h2 -> wf (merge h1 h2)p: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
IHh1: wf h1 -> wf h2 -> disjoint h1 h2 -> wf (merge h1 h2)
H: list_greater p h1
H3: wf h1
H0: wf h2
H1: ~ indom p h2
H2: disjoint h1 h2wf (insert p d (merge h1 h2))p: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
IHh1: wf h1 -> wf h2 -> disjoint h1 h2 -> wf (merge h1 h2)
H: list_greater p h1
H3: wf h1
H0: wf h2
H1: ~ indom p h2
H2: disjoint h1 h2wf (merge h1 h2) -> wf (insert p d (merge h1 h2))p: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
IHh1: wf h1 -> wf h2 -> disjoint h1 h2 -> wf (merge h1 h2)
H: list_greater p h1
H3: wf h1
H0: wf h2
H1: ~ indom p h2
H2: disjoint h1 h2
H4: wf (merge h1 h2)wf (insert p d (merge h1 h2))p: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
IHh1: wf h1 -> wf h2 -> disjoint h1 h2 -> wf (merge h1 h2)
H: list_greater p h1
H3: wf h1
H0: wf h2
H1: ~ indom p h2
H2: disjoint h1 h2
H4: wf (merge h1 h2)fresh p (merge h1 h2)p: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
IHh1: wf h1 -> wf h2 -> disjoint h1 h2 -> wf (merge h1 h2)
H: list_greater p h1
H3: wf h1
H0: wf h2
H1: ~ indom p h2
H2: disjoint h1 h2
H4: wf (merge h1 h2)(list_greater p h1 -> fresh p h1) -> fresh p (merge h1 h2)apply not_indom_merge ; auto. Qed.p: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
IHh1: wf h1 -> wf h2 -> disjoint h1 h2 -> wf (merge h1 h2)
H: list_greater p h1
H3: wf h1
H0: wf h2
H1: ~ indom p h2
H2: disjoint h1 h2
H4: wf (merge h1 h2)
H5: list_greater p h1 -> fresh p h1fresh p (merge h1 h2)
Insertion is commutative for distinct pointers.
h: heap
x1: nat
v1: dynamic
x2: nat
v2: dynamicx1 <> x2 -> insert x1 v1 (insert x2 v2 h) = insert x2 v2 (insert x1 v1 h)h: heap
x1: nat
v1: dynamic
x2: nat
v2: dynamicx1 <> x2 -> insert x1 v1 (insert x2 v2 h) = insert x2 v2 (insert x1 v1 h)x1: nat
v1: dynamic
x2: nat
v2: dynamicx1 <> x2 -> insert x1 v1 (insert x2 v2 nil) = insert x2 v2 (insert x1 v1 nil)a: (ptr * dynamic)%type
h: list (ptr * dynamic)
x1: nat
v1: dynamic
x2: nat
v2: dynamic
IHh: x1 <> x2 -> insert x1 v1 (insert x2 v2 h) = insert x2 v2 (insert x1 v1 h)x1 <> x2 -> insert x1 v1 (insert x2 v2 (a :: h)) = insert x2 v2 (insert x1 v1 (a :: h))a: (ptr * dynamic)%type
h: list (ptr * dynamic)
x1: nat
v1: dynamic
x2: nat
v2: dynamic
IHh: x1 <> x2 -> insert x1 v1 (insert x2 v2 h) = insert x2 v2 (insert x1 v1 h)x1 <> x2 -> insert x1 v1 (insert x2 v2 (a :: h)) = insert x2 v2 (insert x1 v1 (a :: h))p: ptr
d: dynamic
h: list (ptr * dynamic)
x1: nat
v1: dynamic
x2: nat
v2: dynamic
IHh: x1 <> x2 -> insert x1 v1 (insert x2 v2 h) = insert x2 v2 (insert x1 v1 h)x1 <> x2 -> insert x1 v1 (insert x2 v2 ((p, d) :: h)) = insert x2 v2 (insert x1 v1 ((p, d) :: h))p: ptr
d: dynamic
h: list (ptr * dynamic)
x1: nat
v1: dynamic
x2: nat
v2: dynamic
IHh: x1 <> x2 -> insert x1 v1 (insert x2 v2 h) = insert x2 v2 (insert x1 v1 h)
H: x1 <> x2insert x1 v1 (insert x2 v2 ((p, d) :: h)) = insert x2 v2 (insert x1 v1 ((p, d) :: h))p: ptr
d: dynamic
h: list (ptr * dynamic)
x1: nat
v1: dynamic
x2: nat
v2: dynamic
IHh: x1 <> x2 -> insert x1 v1 (insert x2 v2 h) = insert x2 v2 (insert x1 v1 h)
H: x1 <> x2insert x1 v1 (if le_gt_dec x2 p then (x2, v2) :: (p, d) :: h else (p, d) :: insert x2 v2 h) = insert x2 v2 (if le_gt_dec x1 p then (x1, v1) :: (p, d) :: h else (p, d) :: insert x1 v1 h)rewrite IHh ; auto. Qed.p: ptr
d: dynamic
h: list (ptr * dynamic)
x1: nat
v1: dynamic
x2: nat
v2: dynamic
IHh: x1 <> x2 -> insert x1 v1 (insert x2 v2 h) = insert x2 v2 (insert x1 v1 h)
H: x1 <> x2
g: x2 > p
g0: x1 > p
g1: x2 > p(p, d) :: insert x1 v1 (insert x2 v2 h) = (p, d) :: insert x2 v2 (insert x1 v1 h)
Merge is commutative when the heaps are well-formed and disjoint.
h1, h2: heapwf h1 -> wf h2 -> disjoint h1 h2 -> merge h1 h2 = merge h2 h1h1, h2: heapwf h1 -> wf h2 -> disjoint h1 h2 -> merge h1 h2 = merge h2 h1h2: heap
H: True
H0: wf h2
H1: Trueh2 = merge h2 nilp: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
IHh1: wf h1 -> wf h2 -> disjoint h1 h2 -> merge h1 h2 = merge h2 h1
H: list_greater p h1
H3: wf h1
H0: wf h2
H1: ~ indom p h2
H2: disjoint h1 h2insert p d (merge h1 h2) = merge h2 ((p, d) :: h1)p: ptr
d: dynamic
h2: list (ptr * dynamic)
H: True
H0: list_greater p h2
H2: wf h2
H1: True
IHh2: wf h2 -> h2 = merge h2 nil(p, d) :: h2 = insert p d (merge h2 nil)p: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
IHh1: wf h1 -> wf h2 -> disjoint h1 h2 -> merge h1 h2 = merge h2 h1
H: list_greater p h1
H3: wf h1
H0: wf h2
H1: ~ indom p h2
H2: disjoint h1 h2insert p d (merge h1 h2) = merge h2 ((p, d) :: h1)p: ptr
d: dynamic
h2: list (ptr * dynamic)
H: True
H0: list_greater p h2
H2: wf h2
H1: True
IHh2: wf h2 -> h2 = merge h2 nil(p, d) :: h2 = insert p d h2p: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
IHh1: wf h1 -> wf h2 -> disjoint h1 h2 -> merge h1 h2 = merge h2 h1
H: list_greater p h1
H3: wf h1
H0: wf h2
H1: ~ indom p h2
H2: disjoint h1 h2insert p d (merge h1 h2) = merge h2 ((p, d) :: h1)p: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
IHh1: wf h1 -> wf h2 -> disjoint h1 h2 -> merge h1 h2 = merge h2 h1
H: list_greater p h1
H3: wf h1
H0: wf h2
H1: ~ indom p h2
H2: disjoint h1 h2insert p d (merge h1 h2) = merge h2 ((p, d) :: h1)p: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
IHh1: wf h1 -> wf h2 -> disjoint h1 h2 -> merge h1 h2 = merge h2 h1
H: list_greater p h1
H3: wf h1
H0: wf h2
H1: ~ indom p h2
H2: disjoint h1 h2insert p d (merge h2 h1) = merge h2 ((p, d) :: h1)p: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
H: list_greater p h1
H3: wf h1
H0: wf h2
H1: ~ indom p h2
H2: disjoint h1 h2insert p d (merge h2 h1) = merge h2 ((p, d) :: h1)p: ptr
d: dynamic
h1: list (ptr * dynamic)
H: list_greater p h1
H3: wf h1
H0: wf nil
H1: ~ indom p nil
H2: disjoint h1 nilinsert p d (merge nil h1) = merge nil ((p, d) :: h1)p: ptr
d: dynamic
h1: list (ptr * dynamic)
a: (ptr * dynamic)%type
h2: list (ptr * dynamic)
H: list_greater p h1
H3: wf h1
H0: wf (a :: h2)
H1: ~ indom p (a :: h2)
H2: disjoint h1 (a :: h2)
IHh2: wf h2 -> ~ indom p h2 -> disjoint h1 h2 -> insert p d (merge h2 h1) = merge h2 ((p, d) :: h1)insert p d (merge (a :: h2) h1) = merge (a :: h2) ((p, d) :: h1)p: ptr
d: dynamic
h1: list (ptr * dynamic)
H: list_greater p h1
H3: wf h1
H0: wf nil
H1: ~ indom p nil
H2: disjoint h1 nilinsert p d h1 = (p, d) :: h1p: ptr
d: dynamic
h1: list (ptr * dynamic)
a: (ptr * dynamic)%type
h2: list (ptr * dynamic)
H: list_greater p h1
H3: wf h1
H0: wf (a :: h2)
H1: ~ indom p (a :: h2)
H2: disjoint h1 (a :: h2)
IHh2: wf h2 -> ~ indom p h2 -> disjoint h1 h2 -> insert p d (merge h2 h1) = merge h2 ((p, d) :: h1)insert p d (merge (a :: h2) h1) = merge (a :: h2) ((p, d) :: h1)p: ptr
d: dynamic
h1: list (ptr * dynamic)
a: (ptr * dynamic)%type
h2: list (ptr * dynamic)
H: list_greater p h1
H3: wf h1
H0: wf (a :: h2)
H1: ~ indom p (a :: h2)
H2: disjoint h1 (a :: h2)
IHh2: wf h2 -> ~ indom p h2 -> disjoint h1 h2 -> insert p d (merge h2 h1) = merge h2 ((p, d) :: h1)insert p d (merge (a :: h2) h1) = merge (a :: h2) ((p, d) :: h1)p: ptr
d: dynamic
h1: list (ptr * dynamic)
p0: ptr
d0: dynamic
h2: list (ptr * dynamic)
H: list_greater p h1
H3: wf h1
H0: wf ((p0, d0) :: h2)
H1: ~ indom p ((p0, d0) :: h2)
H2: disjoint h1 ((p0, d0) :: h2)
IHh2: wf h2 -> ~ indom p h2 -> disjoint h1 h2 -> insert p d (merge h2 h1) = merge h2 ((p, d) :: h1)insert p d (insert p0 d0 (merge h2 h1)) = insert p0 d0 (merge h2 ((p, d) :: h1))p: ptr
d: dynamic
h1: list (ptr * dynamic)
p0: ptr
d0: dynamic
h2: list (ptr * dynamic)
H: list_greater p h1
H3: wf h1
H0: list_greater p0 h2
H4: wf h2
n: p <> p0
H1: ~ indom p h2
H2: disjoint h1 ((p0, d0) :: h2)
IHh2: wf h2 -> ~ indom p h2 -> disjoint h1 h2 -> insert p d (merge h2 h1) = merge h2 ((p, d) :: h1)insert p d (insert p0 d0 (merge h2 h1)) = insert p0 d0 (merge h2 ((p, d) :: h1))p: ptr
d: dynamic
h1: list (ptr * dynamic)
p0: ptr
d0: dynamic
h2: list (ptr * dynamic)
H: list_greater p h1
H3: wf h1
H0: list_greater p0 h2
H4: wf h2
n: p <> p0
H1: ~ indom p h2
H2: disjoint h1 ((p0, d0) :: h2)
IHh2: wf h2 -> ~ indom p h2 -> disjoint h1 h2 -> insert p d (merge h2 h1) = merge h2 ((p, d) :: h1)insert p d (insert p0 d0 (merge h2 h1)) = insert p0 d0 (insert p d (merge h2 h1))p: ptr
d: dynamic
h1: list (ptr * dynamic)
p0: ptr
d0: dynamic
h2: list (ptr * dynamic)
H: list_greater p h1
H3: wf h1
H0: list_greater p0 h2
H4: wf h2
n: p <> p0
H1: ~ indom p h2
H2: disjoint h1 ((p0, d0) :: h2)
IHh2: wf h2 -> ~ indom p h2 -> disjoint h1 h2 -> insert p d (merge h2 h1) = merge h2 ((p, d) :: h1)disjoint h1 h2p: ptr
d: dynamic
h1: list (ptr * dynamic)
p0: ptr
d0: dynamic
h2: list (ptr * dynamic)
H: list_greater p h1
H3: wf h1
H0: list_greater p0 h2
H4: wf h2
n: p <> p0
H1: ~ indom p h2
H2: disjoint h1 ((p0, d0) :: h2)
IHh2: wf h2 -> ~ indom p h2 -> disjoint h1 h2 -> insert p d (merge h2 h1) = merge h2 ((p, d) :: h1)disjoint h1 h2p: ptr
d: dynamic
h1: list (ptr * dynamic)
p0: ptr
d0: dynamic
h2: list (ptr * dynamic)
H: list_greater p h1
H3: wf h1
H0: list_greater p0 h2
H4: wf h2
n: p <> p0
H1: ~ indom p h2
H2: disjoint h1 ((p0, d0) :: h2)
IHh2: wf h2 -> ~ indom p h2 -> disjoint h1 h2 -> insert p d (merge h2 h1) = merge h2 ((p, d) :: h1)disjoint ((p0, d0) :: h2) h1 -> disjoint h1 h2apply disjoint_comm ; auto. Qed.p: ptr
d: dynamic
h1: list (ptr * dynamic)
p0: ptr
d0: dynamic
h2: list (ptr * dynamic)
H: list_greater p h1
H3: wf h1
H0: list_greater p0 h2
H4: wf h2
n: p <> p0
H1: ~ indom p h2
H2: disjoint h1 ((p0, d0) :: h2)
IHh2: wf h2 -> ~ indom p h2 -> disjoint h1 h2 -> insert p d (merge h2 h1) = merge h2 ((p, d) :: h1)
H5: ~ indom p0 h1
H6: disjoint h2 h1disjoint h1 h2forall h1 h2 h3 : heap, disjoint h1 h3 -> disjoint h2 h3 -> disjoint (merge h1 h2) h3forall h1 h2 h3 : heap, disjoint h1 h3 -> disjoint h2 h3 -> disjoint (merge h1 h2) h3p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, disjoint h1 h3 -> disjoint h2 h3 -> disjoint (merge h1 h2) h3
h2, h3: heap
H: ~ indom p h3
H1: disjoint h1 h3
H0: disjoint h2 h3disjoint (insert p d (merge h1 h2)) h3p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, disjoint h1 h3 -> disjoint h2 h3 -> disjoint (merge h1 h2) h3
h2, h3: heap
H: ~ indom p h3
H1: disjoint h1 h3
H0: disjoint h2 h3disjoint (merge h1 h2) h3 -> disjoint (insert p d (merge h1 h2)) h3induction h ; mysimp. Qed.p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, disjoint h1 h3 -> disjoint h2 h3 -> disjoint (merge h1 h2) h3
h2, h3: heap
H: ~ indom p h3
H1: disjoint h1 h3
H0: disjoint h2 h3forall h : heap, disjoint h h3 -> disjoint (insert p d h) h3
Well-formed merges are associative.
forall h1 h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3forall h1 h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3
h2, h3: heap
H: list_greater p h1
H7: wf h1
H0: wf h2
H1: wf h3
H2: ~ indom p h2
H6: disjoint h1 h2
H3: ~ indom p h3
H5: disjoint h1 h3
H4: disjoint h2 h3insert p d (merge h1 (merge h2 h3)) = merge (insert p d (merge h1 h2)) h3p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3
h2, h3: heap
H: list_greater p h1
H7: wf h1
H0: wf h2
H1: wf h3
H2: ~ indom p h2
H6: disjoint h1 h2
H3: ~ indom p h3
H5: disjoint h1 h3
H4: disjoint h2 h3insert p d (merge (merge h1 h2) h3) = merge (insert p d (merge h1 h2)) h3p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3
h2, h3: heap
H: list_greater p h1
H7: wf h1
H0: wf h2
H1: wf h3
H2: ~ indom p h2
H6: disjoint h1 h2
H3: ~ indom p h3
H5: disjoint h1 h3
H4: disjoint h2 h3wf (merge h1 h2)p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3
h2, h3: heap
H: list_greater p h1
H7: wf h1
H0: wf h2
H1: wf h3
H2: ~ indom p h2
H6: disjoint h1 h2
H3: ~ indom p h3
H5: disjoint h1 h3
H4: disjoint h2 h3
H8: wf (merge h1 h2)insert p d (merge (merge h1 h2) h3) = merge (insert p d (merge h1 h2)) h3p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3
h2, h3: heap
H: list_greater p h1
H7: wf h1
H0: wf h2
H1: wf h3
H2: ~ indom p h2
H6: disjoint h1 h2
H3: ~ indom p h3
H5: disjoint h1 h3
H4: disjoint h2 h3
H8: wf (merge h1 h2)insert p d (merge (merge h1 h2) h3) = merge (insert p d (merge h1 h2)) h3p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3
h2, h3: heap
H: list_greater p h1
H7: wf h1
H0: wf h2
H1: wf h3
H2: ~ indom p h2
H6: disjoint h1 h2
H3: ~ indom p h3
H5: disjoint h1 h3
H4: disjoint h2 h3wf (merge h1 h2) -> insert p d (merge (merge h1 h2) h3) = merge (insert p d (merge h1 h2)) h3p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3
h2, h3: heap
H: list_greater p h1
H7: wf h1
H0: wf h2
H1: wf h3
H2: ~ indom p h2
H6: disjoint h1 h2
H3: ~ indom p h3
H5: disjoint h1 h3
H4: disjoint h2 h3~ indom p (merge h1 h2)p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3
h2, h3: heap
H: list_greater p h1
H7: wf h1
H0: wf h2
H1: wf h3
H2: ~ indom p h2
H6: disjoint h1 h2
H3: ~ indom p h3
H5: disjoint h1 h3
H4: disjoint h2 h3
H8: ~ indom p (merge h1 h2)wf (merge h1 h2) -> insert p d (merge (merge h1 h2) h3) = merge (insert p d (merge h1 h2)) h3p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3
h2, h3: heap
H: list_greater p h1
H7: wf h1
H0: wf h2
H1: wf h3
H2: ~ indom p h2
H6: disjoint h1 h2
H3: ~ indom p h3
H5: disjoint h1 h3
H4: disjoint h2 h3fresh p h1p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3
h2, h3: heap
H: list_greater p h1
H7: wf h1
H0: wf h2
H1: wf h3
H2: ~ indom p h2
H6: disjoint h1 h2
H3: ~ indom p h3
H5: disjoint h1 h3
H4: disjoint h2 h3
H8: ~ indom p (merge h1 h2)wf (merge h1 h2) -> insert p d (merge (merge h1 h2) h3) = merge (insert p d (merge h1 h2)) h3p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3
h2, h3: heap
H: list_greater p h1
H7: wf h1
H0: wf h2
H1: wf h3
H2: ~ indom p h2
H6: disjoint h1 h2
H3: ~ indom p h3
H5: disjoint h1 h3
H4: disjoint h2 h3
H8: ~ indom p (merge h1 h2)wf (merge h1 h2) -> insert p d (merge (merge h1 h2) h3) = merge (insert p d (merge h1 h2)) h3p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3
h2, h3: heap
H: list_greater p h1
H7: wf h1
H0: wf h2
H1: wf h3
H2: ~ indom p h2
H6: disjoint h1 h2
H3: ~ indom p h3
H5: disjoint h1 h3
H4: disjoint h2 h3~ indom p (merge h1 h2) -> wf (merge h1 h2) -> insert p d (merge (merge h1 h2) h3) = merge (insert p d (merge h1 h2)) h3p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3
h2, h3: heap
H: list_greater p h1
H7: wf h1
H0: wf h2
H1: wf h3
H2: ~ indom p h2
H6: disjoint h1 h2
H3: ~ indom p h3
H5: disjoint h1 h3
H4: disjoint h2 h3disjoint (merge h1 h2) h3p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3
h2, h3: heap
H: list_greater p h1
H7: wf h1
H0: wf h2
H1: wf h3
H2: ~ indom p h2
H6: disjoint h1 h2
H3: ~ indom p h3
H5: disjoint h1 h3
H4: disjoint h2 h3
H8: disjoint (merge h1 h2) h3~ indom p (merge h1 h2) -> wf (merge h1 h2) -> insert p d (merge (merge h1 h2) h3) = merge (insert p d (merge h1 h2)) h3p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3
h2, h3: heap
H: list_greater p h1
H7: wf h1
H0: wf h2
H1: wf h3
H2: ~ indom p h2
H6: disjoint h1 h2
H3: ~ indom p h3
H5: disjoint h1 h3
H4: disjoint h2 h3
H8: disjoint (merge h1 h2) h3~ indom p (merge h1 h2) -> wf (merge h1 h2) -> insert p d (merge (merge h1 h2) h3) = merge (insert p d (merge h1 h2)) h3p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3
h2, h3: heap
H: list_greater p h1
H7: wf h1
H0: wf h2
H1: wf h3
H2: ~ indom p h2
H6: disjoint h1 h2
H3: ~ indom p h3
H5: disjoint h1 h3
H4: disjoint h2 h3disjoint (merge h1 h2) h3 -> ~ indom p (merge h1 h2) -> wf (merge h1 h2) -> insert p d (merge (merge h1 h2) h3) = merge (insert p d (merge h1 h2)) h3p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3
h2, h3: heap
H: list_greater p h1
H7: wf h1
H0: wf h2
H1: wf h3
H2: ~ indom p h2
H6: disjoint h1 h2
H3: ~ indom p h3
H5: disjoint h1 h3
H4: disjoint h2 h3forall h : heap, disjoint h h3 -> ~ indom p h -> wf h -> insert p d (merge h h3) = merge (insert p d h) h3p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3
h2, h3: heap
H: list_greater p h1
H7: wf h1
H0: wf h2
H1: wf h3
H2: ~ indom p h2
H6: disjoint h1 h2
H3: ~ indom p h3
H5: disjoint h1 h3
H4: disjoint h2 h3
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: disjoint h h3 -> ~ indom p h -> wf h -> insert p d (merge h h3) = merge (insert p d h) h3
n: p <> p0
g: p > p0
H8: ~ indom p0 h3
H12: disjoint h h3
H9: ~ indom p h
H10: list_greater p0 h
H11: wf hinsert p d (insert p0 d0 (merge h h3)) = insert p0 d0 (merge (insert p d h) h3)rewrite IHh ; auto. Qed.p: ptr
d: dynamic
h1: list (ptr * dynamic)
IHh1: forall h2 h3 : heap, wf h1 -> wf h2 -> wf h3 -> disjoint h1 h2 -> disjoint h1 h3 -> disjoint h2 h3 -> merge h1 (merge h2 h3) = merge (merge h1 h2) h3
h2, h3: heap
H: list_greater p h1
H7: wf h1
H0: wf h2
H1: wf h3
H2: ~ indom p h2
H6: disjoint h1 h2
H3: ~ indom p h3
H5: disjoint h1 h3
H4: disjoint h2 h3
p0: ptr
d0: dynamic
h: list (ptr * dynamic)
IHh: disjoint h h3 -> ~ indom p h -> wf h -> insert p d (merge h h3) = merge (insert p d h) h3
n: p <> p0
g: p > p0
H8: ~ indom p0 h3
H12: disjoint h h3
H9: ~ indom p h
H10: list_greater p0 h
H11: wf hinsert p0 d0 (insert p d (merge h h3)) = insert p0 d0 (merge (insert p d h) h3)x: ptr
h: heap
y: ptr
v: dynamicx <> y -> fresh x (insert y v h) -> fresh x hunfold fresh ; induction h ; simpl in * ; intros ; mysimp ; simpl in * ; mysimp. Qed.x: ptr
h: heap
y: ptr
v: dynamicx <> y -> fresh x (insert y v h) -> fresh x h
If p
is fresh for merge h1 h2
, then it's also fresh for both h1
and h2
.
h1, h2: heap
p: ptrfresh p (merge h1 h2) -> fresh p h1 /\ fresh p h2h1, h2: heap
p: ptrfresh p (merge h1 h2) -> fresh p h1 /\ fresh p h2a: (ptr * dynamic)%type
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p (merge h1 h2) -> ~ indom p h1 /\ ~ indom p h2~ indom p (insert (fst a) (snd a) (merge h1 h2)) -> ~ (let (y, _) := a in if ptr_eq p y then True else indom p h1) /\ ~ indom p h2p0: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p (merge h1 h2) -> ~ indom p h1 /\ ~ indom p h2~ indom p (insert (fst (p0, d)) (snd (p0, d)) (merge h1 h2)) -> ~ (if ptr_eq p p0 then True else indom p h1) /\ ~ indom p h2p0: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p (merge h1 h2) -> ~ indom p h1 /\ ~ indom p h2
H: ~ indom p (insert (fst (p0, d)) (snd (p0, d)) (merge h1 h2))
e: p0 = p~ (if ptr_eq p p0 then True else indom p h1) /\ ~ indom p h2p0: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p (merge h1 h2) -> ~ indom p h1 /\ ~ indom p h2
H: ~ indom p (insert (fst (p0, d)) (snd (p0, d)) (merge h1 h2))
n: p0 <> p~ (if ptr_eq p p0 then True else indom p h1) /\ ~ indom p h2d: dynamic
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p (merge h1 h2) -> ~ indom p h1 /\ ~ indom p h2
H: ~ indom p (insert (fst (p, d)) (snd (p, d)) (merge h1 h2))~ (if ptr_eq p p then True else indom p h1) /\ ~ indom p h2p0: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p (merge h1 h2) -> ~ indom p h1 /\ ~ indom p h2
H: ~ indom p (insert (fst (p0, d)) (snd (p0, d)) (merge h1 h2))
n: p0 <> p~ (if ptr_eq p p0 then True else indom p h1) /\ ~ indom p h2d: dynamic
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p (merge h1 h2) -> ~ indom p h1 /\ ~ indom p h2
H: ~ indom p (insert p d (merge h1 h2))~ (if ptr_eq p p then True else indom p h1) /\ ~ indom p h2p0: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p (merge h1 h2) -> ~ indom p h1 /\ ~ indom p h2
H: ~ indom p (insert (fst (p0, d)) (snd (p0, d)) (merge h1 h2))
n: p0 <> p~ (if ptr_eq p p0 then True else indom p h1) /\ ~ indom p h2d: dynamic
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p (merge h1 h2) -> ~ indom p h1 /\ ~ indom p h2
H: ~ indom p (insert p d (merge h1 h2))Falsed: dynamic
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p (merge h1 h2) -> ~ indom p h1 /\ ~ indom p h2
H: ~ indom p (insert p d (merge h1 h2))
H0: False~ (if ptr_eq p p then True else indom p h1) /\ ~ indom p h2p0: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p (merge h1 h2) -> ~ indom p h1 /\ ~ indom p h2
H: ~ indom p (insert (fst (p0, d)) (snd (p0, d)) (merge h1 h2))
n: p0 <> p~ (if ptr_eq p p0 then True else indom p h1) /\ ~ indom p h2d: dynamic
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p (merge h1 h2) -> ~ indom p h1 /\ ~ indom p h2
H: ~ indom p (insert p d (merge h1 h2))indom p (insert p d (merge h1 h2))d: dynamic
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p (merge h1 h2) -> ~ indom p h1 /\ ~ indom p h2
H: ~ indom p (insert p d (merge h1 h2))
H0: False~ (if ptr_eq p p then True else indom p h1) /\ ~ indom p h2p0: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p (merge h1 h2) -> ~ indom p h1 /\ ~ indom p h2
H: ~ indom p (insert (fst (p0, d)) (snd (p0, d)) (merge h1 h2))
n: p0 <> p~ (if ptr_eq p p0 then True else indom p h1) /\ ~ indom p h2d: dynamic
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p (merge h1 h2) -> ~ indom p h1 /\ ~ indom p h2
H: ~ indom p (insert p d (merge h1 h2))forall h : heap, indom p (insert p d h)d: dynamic
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p (merge h1 h2) -> ~ indom p h1 /\ ~ indom p h2
H: ~ indom p (insert p d (merge h1 h2))
H0: False~ (if ptr_eq p p then True else indom p h1) /\ ~ indom p h2p0: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p (merge h1 h2) -> ~ indom p h1 /\ ~ indom p h2
H: ~ indom p (insert (fst (p0, d)) (snd (p0, d)) (merge h1 h2))
n: p0 <> p~ (if ptr_eq p p0 then True else indom p h1) /\ ~ indom p h2d: dynamic
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p (merge h1 h2) -> ~ indom p h1 /\ ~ indom p h2
H: ~ indom p (insert p d (merge h1 h2))
H0: False~ (if ptr_eq p p then True else indom p h1) /\ ~ indom p h2p0: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p (merge h1 h2) -> ~ indom p h1 /\ ~ indom p h2
H: ~ indom p (insert (fst (p0, d)) (snd (p0, d)) (merge h1 h2))
n: p0 <> p~ (if ptr_eq p p0 then True else indom p h1) /\ ~ indom p h2mysimp; apply IHh1; eapply indom_insert ; eauto. Qed.p0: ptr
d: dynamic
h1: list (ptr * dynamic)
h2: heap
p: ptr
IHh1: ~ indom p (merge h1 h2) -> ~ indom p h1 /\ ~ indom p h2
H: ~ indom p (insert (fst (p0, d)) (snd (p0, d)) (merge h1 h2))
n: p0 <> p~ (if ptr_eq p p0 then True else indom p h1) /\ ~ indom p h2h1, h2, h3: heapdisjoint h1 (merge h2 h3) -> disjoint h1 h2 /\ disjoint h1 h3h1, h2, h3: heapdisjoint h1 (merge h2 h3) -> disjoint h1 h2 /\ disjoint h1 h3h2, h3: heapdisjoint nil (merge h2 h3) -> disjoint nil h2 /\ disjoint nil h3a: (ptr * dynamic)%type
h1: list (ptr * dynamic)
h2, h3: heap
IHh1: disjoint h1 (merge h2 h3) -> disjoint h1 h2 /\ disjoint h1 h3disjoint (a :: h1) (merge h2 h3) -> disjoint (a :: h1) h2 /\ disjoint (a :: h1) h3a: (ptr * dynamic)%type
h1: list (ptr * dynamic)
h2, h3: heap
IHh1: disjoint h1 (merge h2 h3) -> disjoint h1 h2 /\ disjoint h1 h3disjoint (a :: h1) (merge h2 h3) -> disjoint (a :: h1) h2 /\ disjoint (a :: h1) h3a: (ptr * dynamic)%type
h1: list (ptr * dynamic)
h2, h3: heap
IHh1: disjoint h1 (merge h2 h3) -> disjoint h1 h2 /\ disjoint h1 h3(let (x, _) := a in ~ indom x (merge h2 h3) /\ disjoint h1 (merge h2 h3)) -> (let (x, _) := a in ~ indom x h2 /\ disjoint h1 h2) /\ (let (x, _) := a in ~ indom x h3 /\ disjoint h1 h3)p: ptr
d: dynamic
h1: list (ptr * dynamic)
h2, h3: heap
IHh1: disjoint h1 (merge h2 h3) -> disjoint h1 h2 /\ disjoint h1 h3~ indom p (merge h2 h3) /\ disjoint h1 (merge h2 h3) -> (~ indom p h2 /\ disjoint h1 h2) /\ ~ indom p h3 /\ disjoint h1 h3p: ptr
d: dynamic
h1: list (ptr * dynamic)
h2, h3: heap
IHh1: disjoint h1 (merge h2 h3) -> disjoint h1 h2 /\ disjoint h1 h3
H: ~ indom p (merge h2 h3) /\ disjoint h1 (merge h2 h3)(~ indom p h2 /\ disjoint h1 h2) /\ ~ indom p h3 /\ disjoint h1 h3p: ptr
d: dynamic
h1: list (ptr * dynamic)
h2, h3: heap
IHh1: disjoint h1 (merge h2 h3) -> disjoint h1 h2 /\ disjoint h1 h3
H: ~ indom p (merge h2 h3)
H0: disjoint h1 (merge h2 h3)(~ indom p h2 /\ disjoint h1 h2) /\ ~ indom p h3 /\ disjoint h1 h3p: ptr
d: dynamic
h1: list (ptr * dynamic)
h2, h3: heap
IHh1: disjoint h1 (merge h2 h3) -> disjoint h1 h2 /\ disjoint h1 h3
H: ~ indom p (merge h2 h3)
H0: disjoint h1 (merge h2 h3)
H1: fresh p h2
H2: fresh p h3(~ indom p h2 /\ disjoint h1 h2) /\ ~ indom p h3 /\ disjoint h1 h3mysimp. Qed.p: ptr
d: dynamic
h1: list (ptr * dynamic)
h2, h3: heap
IHh1: disjoint h1 (merge h2 h3) -> disjoint h1 h2 /\ disjoint h1 h3
H: ~ indom p (merge h2 h3)
H0: disjoint h1 (merge h2 h3)
H1: fresh p h2
H2: fresh p h3disjoint h1 h2 /\ disjoint h1 h3 -> (~ indom p h2 /\ disjoint h1 h2) /\ ~ indom p h3 /\ disjoint h1 h3forall (h : heap) (n : nat), n > max_heap h -> fresh n hforall (h : heap) (n : nat), n > max_heap h -> fresh n hp: ptr
d: dynamic
h: list (ptr * dynamic)
IHh: forall n : nat, n > max_heap h -> ~ indom n h
H: p > (if le_gt_dec p (max_heap h) then max_heap h else p)~ Truep: ptr
d: dynamic
h: list (ptr * dynamic)
IHh: forall n : nat, n > max_heap h -> ~ indom n h
n: nat
H: n > (if le_gt_dec p (max_heap h) then max_heap h else p)
n0: n <> p~ indom n hp: ptr
d: dynamic
h: list (ptr * dynamic)
IHh: forall n : nat, n > max_heap h -> ~ indom n h
n: nat
H: n > (if le_gt_dec p (max_heap h) then max_heap h else p)
n0: n <> p~ indom n hp: ptr
d: dynamic
h: list (ptr * dynamic)
IHh: forall n : nat, n > max_heap h -> ~ indom n h
n: nat
g: p > max_heap h
H: n > p
n0: n <> p~ indom n hp: ptr
d: dynamic
h: list (ptr * dynamic)
IHh: forall n : nat, n > max_heap h -> ~ indom n h
n: nat
g: p > max_heap h
H: n > p
n0: n <> pn > max_heap hlia. Qed.p: ptr
d: dynamic
h: list (ptr * dynamic)
IHh: forall n : nat, n > max_heap h -> ~ indom n h
n: nat
g: p > max_heap h
H: n > p
n0: n <> pn > max_heap hforall (p : ptr) (h : heap), wf h -> fresh p (remove p h)forall (p : ptr) (h : heap), wf h -> fresh p (remove p h)apply lg_imp_not_indom ; auto. Qed.p0: ptr
d: dynamic
h: list (ptr * dynamic)
IHh: wf h -> ~ indom p0 (remove p0 h)
H: list_greater p0 h
H0: wf h~ indom p0 hforall (x : ptr) (v : dynamic) (h : heap), lookup x (insert x v h) = Some vforall (x : ptr) (v : dynamic) (h : heap), lookup x (insert x v h) = Some vv: dynamic
p: ptr
d: dynamic
h: list (ptr * dynamic)
g: p > p
IHh: lookup p (insert p v h) = Some vSome d = Some vv: dynamic
p: ptr
d: dynamic
h: list (ptr * dynamic)
g: p > p
IHh: lookup p (insert p v h) = Some vFalsev: dynamic
p: ptr
d: dynamic
h: list (ptr * dynamic)
g: p > p
IHh: lookup p (insert p v h) = Some v
H: FalseSome d = Some vcontradiction. Qed.v: dynamic
p: ptr
d: dynamic
h: list (ptr * dynamic)
g: p > p
IHh: lookup p (insert p v h) = Some v
H: FalseSome d = Some vforall (x : ptr) (v : dynamic) (h : heap), remove x (insert x v h) = hforall (x : ptr) (v : dynamic) (h : heap), remove x (insert x v h) = hv: dynamic
p: ptr
d: dynamic
h: list (ptr * dynamic)
g: p > p
IHh: remove p (insert p v h) = hinsert p v h = (p, d) :: hv: dynamic
p: ptr
d: dynamic
h: list (ptr * dynamic)
g: p > p
IHh: remove p (insert p v h) = hFalsev: dynamic
p: ptr
d: dynamic
h: list (ptr * dynamic)
g: p > p
IHh: remove p (insert p v h) = h
H: Falseinsert p v h = (p, d) :: hcontradiction. Qed.v: dynamic
p: ptr
d: dynamic
h: list (ptr * dynamic)
g: p > p
IHh: remove p (insert p v h) = h
H: Falseinsert p v h = (p, d) :: h
Reasoning directly in terms of heaps is painful. So we will define some rules for reasoning directly at the level of the separation logic. This list of lemmas is by no means complete...
P: hpropP ==> emp ** PP: hpropP ==> emp ** PP: hprop
h: heap
H: wf h
H0: P hexists h1 h2 : heap, wf h1 /\ wf h2 /\ h1 = nil /\ P h2 /\ h = merge h1 h2 /\ (fix disjoint (h0 h3 : heap) {struct h0} : Prop := match h0 with | nil => True | (x, _) :: h1' => ~ indom x h3 /\ disjoint h1' h3 end) h1 h2P: hprop
h: heap
H: wf h
H0: P hexists h2 : heap, wf empty_heap /\ wf h2 /\ empty_heap = nil /\ P h2 /\ h = merge empty_heap h2 /\ (fix disjoint (h1 h0 : heap) {struct h1} : Prop := match h1 with | nil => True | (x, _) :: h1' => ~ indom x h0 /\ disjoint h1' h0 end) empty_heap h2mysimp. Qed. Hint Resolve emp_star : sep_db.P: hprop
h: heap
H: wf h
H0: P hwf empty_heap /\ wf h /\ empty_heap = nil /\ P h /\ h = merge empty_heap h /\ (fix disjoint (h1 h2 : heap) {struct h1} : Prop := match h1 with | nil => True | (x, _) :: h1' => ~ indom x h2 /\ disjoint h1' h2 end) empty_heap hP, Q: hpropP ** Q ==> Q ** Punfold himp, star ; mysimp ; exists x0 ; exists x ; simpl in * ; mysimp ; try rewrite merge_comm ; auto ; apply disjoint_comm ; simpl ; auto. Qed.P, Q: hpropP ** Q ==> Q ** PP: hpropP ==> P ** empP: hpropP ==> P ** empP: hprop
h: heap
H: wf h
H0: P hexists h1 h2 : heap, wf h1 /\ wf h2 /\ P h1 /\ h2 = nil /\ h = merge h1 h2 /\ disjoint h1 h2apply emp_star ; auto. Qed. Hint Resolve star_emp : sep_db.P: hprop
h: heap
H: wf h
H0: P h((fun h2 : heap => h2 = nil) ** P) hP, Q, R: hpropP ** Q ** R ==> (P ** Q) ** RP, Q, R: hpropP ** Q ** R ==> (P ** Q) ** RP, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)exists h1 h2 : heap, wf h1 /\ wf h2 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ P h3 /\ Q h4 /\ h1 = merge h3 h4 /\ disjoint h3 h4) /\ R h2 /\ merge x (merge x1 x2) = merge h1 h2 /\ disjoint h1 h2P, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)disjoint x x1 /\ disjoint x x2 -> exists h1 h2 : heap, wf h1 /\ wf h2 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ P h3 /\ Q h4 /\ h1 = merge h3 h4 /\ disjoint h3 h4) /\ R h2 /\ merge x (merge x1 x2) = merge h1 h2 /\ disjoint h1 h2P, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)
H4: disjoint x x1
H9: disjoint x x2exists h1 h2 : heap, wf h1 /\ wf h2 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ P h3 /\ Q h4 /\ h1 = merge h3 h4 /\ disjoint h3 h4) /\ R h2 /\ merge x (merge x1 x2) = merge h1 h2 /\ disjoint h1 h2P, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)
H4: disjoint x x1
H9: disjoint x x2exists h2 : heap, wf (merge x x1) /\ wf h2 /\ (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P h1 /\ Q h3 /\ merge x x1 = merge h1 h3 /\ disjoint h1 h3) /\ R h2 /\ merge x (merge x1 x2) = merge (merge x x1) h2 /\ disjoint (merge x x1) h2P, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)
H4: disjoint x x1
H9: disjoint x x2wf (merge x x1) /\ wf x2 /\ (exists h1 h2 : heap, wf h1 /\ wf h2 /\ P h1 /\ Q h2 /\ merge x x1 = merge h1 h2 /\ disjoint h1 h2) /\ R x2 /\ merge x (merge x1 x2) = merge (merge x x1) x2 /\ disjoint (merge x x1) x2P, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)
H4: disjoint x x1
H9: disjoint x x2wf (merge x x1)P, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)
H4: disjoint x x1
H9: disjoint x x2exists h1 h2 : heap, wf h1 /\ wf h2 /\ P h1 /\ Q h2 /\ merge x x1 = merge h1 h2 /\ disjoint h1 h2P, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)
H4: disjoint x x1
H9: disjoint x x2merge x (merge x1 x2) = merge (merge x x1) x2P, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)
H4: disjoint x x1
H9: disjoint x x2disjoint (merge x x1) x2P, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)
H4: disjoint x x1
H9: disjoint x x2exists h1 h2 : heap, wf h1 /\ wf h2 /\ P h1 /\ Q h2 /\ merge x x1 = merge h1 h2 /\ disjoint h1 h2P, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)
H4: disjoint x x1
H9: disjoint x x2merge x (merge x1 x2) = merge (merge x x1) x2P, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)
H4: disjoint x x1
H9: disjoint x x2disjoint (merge x x1) x2P, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)
H4: disjoint x x1
H9: disjoint x x2exists h2 : heap, wf x /\ wf h2 /\ P x /\ Q h2 /\ merge x x1 = merge x h2 /\ disjoint x h2P, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)
H4: disjoint x x1
H9: disjoint x x2merge x (merge x1 x2) = merge (merge x x1) x2P, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)
H4: disjoint x x1
H9: disjoint x x2disjoint (merge x x1) x2P, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)
H4: disjoint x x1
H9: disjoint x x2wf x /\ wf x1 /\ P x /\ Q x1 /\ merge x x1 = merge x x1 /\ disjoint x x1P, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)
H4: disjoint x x1
H9: disjoint x x2merge x (merge x1 x2) = merge (merge x x1) x2P, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)
H4: disjoint x x1
H9: disjoint x x2disjoint (merge x x1) x2P, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)
H4: disjoint x x1
H9: disjoint x x2merge x (merge x1 x2) = merge (merge x x1) x2P, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)
H4: disjoint x x1
H9: disjoint x x2disjoint (merge x x1) x2apply disjoint_merge ; auto. Qed.P, Q, R: hprop
x, x1, x2: heap
H: wf (merge x (merge x1 x2))
H0: wf x
H1: wf (merge x1 x2)
H2: P x
H3: wf x1
H6: wf x2
H7: Q x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint x (merge x1 x2)
H4: disjoint x x1
H9: disjoint x x2disjoint (merge x x1) x2P, R: hprop
Q: PropQ -> R ==> P -> R ==> pure Q ** PP, R: hprop
Q: PropQ -> R ==> P -> R ==> pure Q ** PP, R: hprop
Q: PropQ -> (forall h : heap, wf h -> R h -> P h) -> forall h : heap, wf h -> R h -> exists h1 h2 : heap, wf h1 /\ wf h2 /\ (h1 = nil /\ Q) /\ P h2 /\ h = merge h1 h2 /\ (fix disjoint (h0 h3 : heap) {struct h0} : Prop := match h0 with | nil => True | (x, _) :: h1' => ~ indom x h3 /\ disjoint h1' h3 end) h1 h2P, R: hprop
Q: Prop
H: Q
H0: forall h : heap, wf h -> R h -> P h
h: heap
H1: wf h
H2: R hexists h1 h2 : heap, wf h1 /\ wf h2 /\ (h1 = nil /\ Q) /\ P h2 /\ h = merge h1 h2 /\ (fix disjoint (h0 h3 : heap) {struct h0} : Prop := match h0 with | nil => True | (x, _) :: h1' => ~ indom x h3 /\ disjoint h1' h3 end) h1 h2P, R: hprop
Q: Prop
H: Q
H0: forall h : heap, wf h -> R h -> P h
h: heap
H1: wf h
H2: R hexists h2 : heap, wf empty_heap /\ wf h2 /\ (empty_heap = nil /\ Q) /\ P h2 /\ h = merge empty_heap h2 /\ (fix disjoint (h1 h0 : heap) {struct h1} : Prop := match h1 with | nil => True | (x, _) :: h1' => ~ indom x h0 /\ disjoint h1' h0 end) empty_heap h2mysimp. Qed. Hint Resolve pure_elim : sep_db.P, R: hprop
Q: Prop
H: Q
H0: forall h : heap, wf h -> R h -> P h
h: heap
H1: wf h
H2: R hwf empty_heap /\ wf h /\ (empty_heap = nil /\ Q) /\ P h /\ h = merge empty_heap h /\ (fix disjoint (h1 h2 : heap) {struct h1} : Prop := match h1 with | nil => True | (x, _) :: h1' => ~ indom x h2 /\ disjoint h1' h2 end) empty_heap hP, Q, R: hpropP ==> Q ** R -> P ==> R ** QP, Q, R: hpropP ==> Q ** R -> P ==> R ** QP, Q, R: hprop
H: P ==> Q ** RP ==> R ** Qapply star_comm ; auto. Qed.P, Q, R: hprop
H: P ==> Q ** R
h: heap
Hwf: wf h
HP: P h(R ** Q) hforall P : hprop, P ==> Pauto. Qed. Hint Resolve himp_id : sep_db.forall (P : hprop) (h : heap), wf h -> P h -> P hforall (t : stype) (x : tptr t) (v : interp t), x --> v ==> x -->?forall (t : stype) (x : tptr t) (v : interp t) (h : heap), wf h -> h = (x, existT interp t v) :: nil -> exists x0 : dynamic, h = (x, x0) :: nilt: stype
x: tptr t
v: interp t
h: heap
H: wf h
H0: h = (x, existT interp t v) :: nilexists x0 : dynamic, h = (x, x0) :: nilauto. Qed. Hint Resolve ptsto_ptsto_some : sep_db.t: stype
x: tptr t
v: interp t
h: heap
H: wf h
H0: h = (x, existT interp t v) :: nilh = (x, existT interp t v) :: nilforall P Q R : hprop, P ** Q ==> R -> Q ** P ==> Rforall P Q R : hprop, P ** Q ==> R -> Q ** P ==> RP, Q, R: hprop
H: forall h : heap, wf h -> (P ** Q) h -> R h
h: heap
H0: wf h
H1: (Q ** P) hR happly star_comm ; auto. Qed.P, Q, R: hprop
H: forall h : heap, wf h -> (P ** Q) h -> R h
h: heap
H0: wf h
H1: (Q ** P) h(P ** Q) hforall (P : Prop) (Q R : hprop), (P -> Q ==> R) -> pure P ** Q ==> Rforall (P : Prop) (Q R : hprop), (P -> Q ==> R) -> pure P ** Q ==> Rapply H ; auto. Qed. Hint Resolve hyp_pure : sep_db.P: Prop
Q, R: hprop
H: P -> forall h : heap, wf h -> Q h -> R h
x0: heap
H0: wf x0
H1: True
H2: wf x0
H7: P
H4: Q x0
H6: TrueR x0forall (t : stype) (p : tptr t) (v : interp t), p --> v ==> hexists (fun v0 : interp t => p --> v0)forall (t : stype) (p : tptr t) (v : interp t), p --> v ==> hexists (fun v0 : interp t => p --> v0)eauto. Qed. Hint Resolve ptsto_hexist : sep_db.t: stype
p: tptr t
v: interp t
H: wf ((p, existT interp t v) :: nil)exists x : interp t, (p, existT interp t v) :: nil = (p, existT interp t x) :: nilforall (t : Type) (F : t -> hprop) (R : hprop), (forall x : t, F x ==> R) -> hexists F ==> Rforall (t : Type) (F : t -> hprop) (R : hprop), (forall x : t, F x ==> R) -> hexists F ==> Rforall (t : Type) (F : t -> hprop) (R : hprop), (forall x : t, F x ==> R) -> (fun h : heap => exists x : t, F x h) ==> Rt: Type
F: t -> hprop
R: hprop
H: forall x : t, F x ==> R(fun h : heap => exists x : t, F x h) ==> Rt: Type
F: t -> hprop
R: hprop
H: forall x : t, F x ==> R
h: heap
Hwf: wf h(exists x : t, F x h) -> R happly (H x h) ; auto. Qed.t: Type
F: t -> hprop
R: hprop
H: forall x : t, F x ==> R
h: heap
Hwf: wf h
x: t
H0: F x hR hP1, P2, Q1, Q2: hpropP1 ==> Q1 -> P2 ==> Q2 -> P1 ** P2 ==> Q1 ** Q2P1, P2, Q1, Q2: hpropP1 ==> Q1 -> P2 ==> Q2 -> P1 ** P2 ==> Q1 ** Q2P1, P2, Q1, Q2: hprop
H: forall h : heap, wf h -> P1 h -> Q1 h
H0: forall h : heap, wf h -> P2 h -> Q2 h
h: heap
H1: wf h
x, x0: heap
H2: wf x
H3: wf x0
H4: P1 x
H5: P2 x0
H6: h = merge x x0
H7: disjoint x x0exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q1 h1 /\ Q2 h2 /\ h = merge h1 h2 /\ disjoint h1 h2P1, P2, Q1, Q2: hprop
H: forall h : heap, wf h -> P1 h -> Q1 h
H0: forall h : heap, wf h -> P2 h -> Q2 h
h: heap
H1: wf h
x, x0: heap
H2: wf x
H3: wf x0
H4: P1 x
H5: P2 x0
H6: h = merge x x0
H7: disjoint x x0exists h2 : heap, wf x /\ wf h2 /\ Q1 x /\ Q2 h2 /\ h = merge x h2 /\ disjoint x h2mysimp. Qed.P1, P2, Q1, Q2: hprop
H: forall h : heap, wf h -> P1 h -> Q1 h
H0: forall h : heap, wf h -> P2 h -> Q2 h
h: heap
H1: wf h
x, x0: heap
H2: wf x
H3: wf x0
H4: P1 x
H5: P2 x0
H6: h = merge x x0
H7: disjoint x x0wf x /\ wf x0 /\ Q1 x /\ Q2 x0 /\ h = merge x x0 /\ disjoint x x0
This little tactic helps prove things are disjoint:
Ltac disj := repeat match goal with | [ H : disjoint ?x ?y |- disjoint ?y ?x ] => apply disjoint_comm ; auto | [ |- disjoint (merge _ _) _ ] => apply disjoint_merge | [ |- disjoint _ (merge _ _) ] => apply disjoint_comm | [ H : disjoint (merge _ _) _ |- _ ] => generalize (merge_disjoint _ _ _ (disjoint_comm _ _ H)) ; clear H ; mysimp | [ H : disjoint _ (merge _ _) |- _ ] => generalize (merge_disjoint _ _ _ H) ; clear H ; mysimp | [ |- wf (merge _ _) ] => apply merge_wf ; mysimp | [ |- _ ] => assumption end.P, Q, R: hprop(P ** Q) ** R ==> P ** Q ** RP, Q, R: hprop(P ** Q) ** R ==> P ** Q ** RP, Q, R: hpropforall h : heap, wf h -> (exists h1 h2 : heap, wf h1 /\ wf h2 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ P h3 /\ Q h4 /\ h1 = merge h3 h4 /\ disjoint h3 h4) /\ R h2 /\ h = merge h1 h2 /\ disjoint h1 h2) -> exists h1 h2 : heap, wf h1 /\ wf h2 /\ P h1 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ Q h3 /\ R h4 /\ h2 = merge h3 h4 /\ disjoint h3 h4) /\ h = merge h1 h2 /\ disjoint h1 h2P, Q, R: hprop
h: heap
H: wf h
x, x0: heap
H0: wf x
H1: wf x0
x1, x2: heap
H2: wf x1
H6: wf x2
H7: P x1
H8: Q x2
H9: x = merge x1 x2
H10: disjoint x1 x2
H3: R x0
H4: h = merge x x0
H5: disjoint x x0exists h1 h2 : heap, wf h1 /\ wf h2 /\ P h1 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ Q h3 /\ R h4 /\ h2 = merge h3 h4 /\ disjoint h3 h4) /\ h = merge h1 h2 /\ disjoint h1 h2P, Q, R: hprop
x0, x1, x2: heap
H: wf (merge (merge x1 x2) x0)
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: Q x2
H10: disjoint x1 x2
H3: R x0
H5: disjoint (merge x1 x2) x0exists h1 h2 : heap, wf h1 /\ wf h2 /\ P h1 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ Q h3 /\ R h4 /\ h2 = merge h3 h4 /\ disjoint h3 h4) /\ merge (merge x1 x2) x0 = merge h1 h2 /\ disjoint h1 h2P, Q, R: hprop
x0, x1, x2: heap
H: wf (merge (merge x1 x2) x0)
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: Q x2
H10: disjoint x1 x2
H3: R x0
H5: disjoint (merge x1 x2) x0exists h2 : heap, wf x1 /\ wf h2 /\ P x1 /\ (exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q h1 /\ R h3 /\ h2 = merge h1 h3 /\ disjoint h1 h3) /\ merge (merge x1 x2) x0 = merge x1 h2 /\ disjoint x1 h2P, Q, R: hprop
x0, x1, x2: heap
H: wf (merge (merge x1 x2) x0)
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: Q x2
H10: disjoint x1 x2
H3: R x0
H5: disjoint (merge x1 x2) x0wf x1 /\ wf (merge x2 x0) /\ P x1 /\ (exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q h1 /\ R h2 /\ merge x2 x0 = merge h1 h2 /\ disjoint h1 h2) /\ merge (merge x1 x2) x0 = merge x1 (merge x2 x0) /\ disjoint x1 (merge x2 x0)P, Q, R: hprop
x0, x1, x2: heap
H: wf (merge (merge x1 x2) x0)
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: Q x2
H10: disjoint x1 x2
H3: R x0
H4: disjoint x0 x1
H5: disjoint x0 x2exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q h1 /\ R h2 /\ merge x2 x0 = merge h1 h2 /\ disjoint h1 h2P, Q, R: hprop
x0, x1, x2: heap
H: wf (merge (merge x1 x2) x0)
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: Q x2
H10: disjoint x1 x2
H3: R x0
H4: disjoint x0 x1
H5: disjoint x0 x2merge (merge x1 x2) x0 = merge x1 (merge x2 x0)P, Q, R: hprop
x0, x1, x2: heap
H: wf (merge (merge x1 x2) x0)
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: Q x2
H10: disjoint x1 x2
H3: R x0
H4: disjoint x0 x1
H5: disjoint x0 x2exists h2 : heap, wf x2 /\ wf h2 /\ Q x2 /\ R h2 /\ merge x2 x0 = merge x2 h2 /\ disjoint x2 h2P, Q, R: hprop
x0, x1, x2: heap
H: wf (merge (merge x1 x2) x0)
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: Q x2
H10: disjoint x1 x2
H3: R x0
H4: disjoint x0 x1
H5: disjoint x0 x2merge (merge x1 x2) x0 = merge x1 (merge x2 x0)P, Q, R: hprop
x0, x1, x2: heap
H: wf (merge (merge x1 x2) x0)
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: Q x2
H10: disjoint x1 x2
H3: R x0
H4: disjoint x0 x1
H5: disjoint x0 x2wf x2 /\ wf x0 /\ Q x2 /\ R x0 /\ merge x2 x0 = merge x2 x0 /\ disjoint x2 x0P, Q, R: hprop
x0, x1, x2: heap
H: wf (merge (merge x1 x2) x0)
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: Q x2
H10: disjoint x1 x2
H3: R x0
H4: disjoint x0 x1
H5: disjoint x0 x2merge (merge x1 x2) x0 = merge x1 (merge x2 x0)rewrite merge_assoc ; auto ; disj. Qed.P, Q, R: hprop
x0, x1, x2: heap
H: wf (merge (merge x1 x2) x0)
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: Q x2
H10: disjoint x1 x2
H3: R x0
H4: disjoint x0 x1
H5: disjoint x0 x2merge (merge x1 x2) x0 = merge x1 (merge x2 x0)
This lemma will help me simplify some reasoning involving existentials.
A: Type
F: A -> Prop
Y, Z: Prop(exists x : A, F x) -> (Y -> Z) -> ((exists x : A, F x) -> Y) -> ZA: Type
F: A -> Prop
Y, Z: Prop
x: A
H: F x
H0: Y -> Z
H1: (exists x : A, F x) -> YZA: Type
F: A -> Prop
Y, Z: Prop
x: A
H: F x
H0: Y -> Z
H1: (exists x : A, F x) -> YYeauto. Qed.A: Type
F: A -> Prop
Y, Z: Prop
x: A
H: F x
H0: Y -> Z
H1: (exists x : A, F x) -> Yexists x : A, F xP1, P2, P3, Q: hpropP1 ** P2 ** P3 ==> Q -> (P1 ** P2) ** P3 ==> QP1, P2, P3, Q: hpropP1 ** P2 ** P3 ==> Q -> (P1 ** P2) ** P3 ==> QP1, P2, P3, Q: hprop
H: forall h : heap, wf h -> (P1 ** P2 ** P3) h -> Q h
h: heap
H0: wf h
H1: ((P1 ** P2) ** P3) hQ happly star_assoc_l ; auto. Qed. Hint Resolve hyp_assoc_r : sep_db.P1, P2, P3, Q: hprop
H: forall h : heap, wf h -> (P1 ** P2 ** P3) h -> Q h
h: heap
H0: wf h
H1: ((P1 ** P2) ** P3) h(P1 ** P2 ** P3) hP, Q: hpropP ==> Q -> P ==> Q ** empP, Q: hprop(forall h : heap, wf h -> P h -> Q h) -> forall h : heap, wf h -> P h -> exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q h1 /\ h2 = nil /\ h = merge h1 h2 /\ disjoint h1 h2P, Q: hprop
H: forall h : heap, wf h -> P h -> Q h
h: heap
H0: wf h
H1: P hexists h1 h2 : heap, wf h1 /\ wf h2 /\ Q h1 /\ h2 = nil /\ h = merge h1 h2 /\ disjoint h1 h2P, Q: hprop
H: forall h : heap, wf h -> P h -> Q h
h: heap
H0: wf h
H1: P h((fun h2 : heap => h2 = nil) ** Q) hP, Q: hprop
H: forall h : heap, wf h -> P h -> Q h
h: heap
H0: wf h
H1: P hexists h1 h2 : heap, wf h1 /\ wf h2 /\ h1 = nil /\ Q h2 /\ h = merge h1 h2 /\ disjoint h1 h2P, Q: hprop
H: forall h : heap, wf h -> P h -> Q h
h: heap
H0: wf h
H1: P hexists h2 : heap, wf nil /\ wf h2 /\ nil = nil /\ Q h2 /\ h = merge nil h2 /\ disjoint nil h2mysimp. Qed. Hint Resolve conc_emp : sep_db.P, Q: hprop
H: forall h : heap, wf h -> P h -> Q h
h: heap
H0: wf h
H1: P hwf nil /\ wf h /\ nil = nil /\ Q h /\ h = merge nil h /\ disjoint nil hP, Q: hpropP ==> Q -> emp ** P ==> QP, Q: hprop(forall h : heap, wf h -> P h -> Q h) -> forall h : heap, wf h -> (exists h1 h2 : heap, wf h1 /\ wf h2 /\ h1 = nil /\ P h2 /\ h = merge h1 h2 /\ disjoint h1 h2) -> Q hP, Q: hprop
H: forall h : heap, wf h -> P h -> Q h
h: heap
H0: wf h
x, x0: heap
H1: wf x
H2: wf x0
H3: x = nil
H4: P x0
H5: h = merge x x0
H6: disjoint x x0Q hP, Q: hprop
H: forall h : heap, wf h -> P h -> Q h
x0: heap
H0: wf (merge nil x0)
H1: wf nil
H2: wf x0
H4: P x0
H6: disjoint nil x0Q (merge nil x0)auto. Qed. Hint Resolve hyp_emp : sep_db.P, Q: hprop
H: forall h : heap, wf h -> P h -> Q h
x0: heap
H0: wf x0
H1: True
H2: wf x0
H4: P x0
H6: TrueQ x0P, Q1, Q2, Q3: hpropP ==> Q1 ** Q2 ** Q3 -> P ==> (Q1 ** Q2) ** Q3P, Q1, Q2, Q3: hpropP ==> Q1 ** Q2 ** Q3 -> P ==> (Q1 ** Q2) ** Q3apply star_assoc ; auto. Qed.P, Q1, Q2, Q3: hprop
H: forall h : heap, wf h -> P h -> (Q1 ** Q2 ** Q3) h
h: heap
H0: wf h
H1: P h((Q1 ** Q2) ** Q3) hQ, P, R: hpropQ ==> R -> P ==> Q -> P ==> Runfold himp ; mysimp. Qed.Q, P, R: hpropQ ==> R -> P ==> Q -> P ==> R
In separation logic, a total correctness assertion:
[{{ P }} c {{ Q }}]
holds iff (1) we start with a heap h
that can be broken into two parts, one that satisfies P
and another satisfying sing h2
for some h2, (2) we run the command c
on heap h
and get out Some
heap h'
and value v
, and (3) the output heap h'
satisfies Q v ** sing h2
. That is, the output heap can be broken into two disjoint heaps, one satisfying Q v
, and one satisfying sing h2
.
This effectively forces the command to be parametric in some part of the heap (h2
) and leave it alone. In turn, this means that if we have proven a separate property about h2
, this property will be preserved when we run c
.
Definition sep_tc_triple(t:Type) (P:hprop)(c:Cmd t)(Q:t -> hprop) := forall h h2, (P ** sing h2) h -> match c h with | None => False | Some (h',v) => (Q v ** sing h2) h' end. Notation "{{ P }} c {{ Q }}" := (sep_tc_triple P c Q) (at level 90) : cmd_scope.
Lots of definitions to be unwound...
Ltac unf := unfold sep_tc_triple, star, hexists, pure, emp, sing.
This says that ret v
can be run in a heap satisfying emp
and returns a heap satisfying emp
and the return value is equal to v. But remember that this really means that ret
can be run in any heap h
that can be broken into a portion satisfying emp
(i.e., the empty heap and some other heap (which must be h
!) and that the other portion will be preserved. In short, the specification captures the fact that ret
will not change the heap.
t: Type
v: t{{emp}} ret v {{fun x : t => pure (x = v)}}t: Type
v: t{{emp}} ret v {{fun x : t => pure (x = v)}}t: Type
v: t
x0: heap
H: wf nil
H0: wf x0
H4: disjoint nil x0exists h1 h2 : heap, wf h1 /\ wf h2 /\ (h1 = nil /\ v = v) /\ x0 = h2 /\ merge nil x0 = merge h1 h2 /\ disjoint h1 h2t: Type
v: t
x0: heap
H: wf nil
H0: wf x0
H4: disjoint nil x0exists h2 : heap, wf empty_heap /\ wf h2 /\ (empty_heap = nil /\ v = v) /\ x0 = h2 /\ merge nil x0 = merge empty_heap h2 /\ disjoint empty_heap h2mysimp. Qed.t: Type
v: t
x0: heap
H: wf nil
H0: wf x0
H4: disjoint nil x0wf empty_heap /\ wf x0 /\ (empty_heap = nil /\ v = v) /\ x0 = x0 /\ merge nil x0 = merge empty_heap x0 /\ disjoint empty_heap x0
new v
consumes an empty heap, and produces a pointer x
and a heap satisfying x --> v
. Because of the definition of the separation-total-correctness triple, x
must be fresh for the whole heap.
t: stype
v: interp t{{emp}} new t v {{fun x : tptr t => x --> v}}t: stype
v: interp t{{emp}} new t v {{fun x : tptr t => x --> v}}t: stype
v: interp t
x0: heap
H: wf nil
H0: wf x0
H4: disjoint nil x0exists h1 h2 : heap, wf h1 /\ wf h2 /\ (S (max_heap (merge nil x0)) ---> existT interp t v) h1 /\ x0 = h2 /\ insert (S (max_heap (merge nil x0))) (existT interp t v) (merge nil x0) = merge h1 h2 /\ disjoint h1 h2t: stype
v: interp t
x0: heap
H: wf nil
H0: wf x0
H4: disjoint nil x0exists h2 : heap, wf ((S (max_heap x0), existT interp t v) :: nil) /\ wf h2 /\ (S (max_heap (merge nil x0)) ---> existT interp t v) ((S (max_heap x0), existT interp t v) :: nil) /\ x0 = h2 /\ insert (S (max_heap (merge nil x0))) (existT interp t v) (merge nil x0) = merge ((S (max_heap x0), existT interp t v) :: nil) h2 /\ disjoint ((S (max_heap x0), existT interp t v) :: nil) h2t: stype
v: interp t
x0: heap
H: wf nil
H0: wf x0
H4: disjoint nil x0wf ((S (max_heap x0), existT interp t v) :: nil) /\ wf x0 /\ (S (max_heap (merge nil x0)) ---> existT interp t v) ((S (max_heap x0), existT interp t v) :: nil) /\ x0 = x0 /\ insert (S (max_heap (merge nil x0))) (existT interp t v) (merge nil x0) = merge ((S (max_heap x0), existT interp t v) :: nil) x0 /\ disjoint ((S (max_heap x0), existT interp t v) :: nil) x0t: stype
v: interp t
x0: heap
H: wf nil
H0: wf x0
H4: disjoint nil x0(S (max_heap x0) ---> existT interp t v) ((S (max_heap x0), existT interp t v) :: nil)t: stype
v: interp t
x0: heap
H: wf nil
H0: wf x0
H4: disjoint nil x0~ indom (S (max_heap x0)) x0t: stype
v: interp t
x0: heap
H: wf nil
H0: wf x0
H4: disjoint nil x0~ indom (S (max_heap x0)) x0lia. Qed.t: stype
v: interp t
x0: heap
H: wf nil
H0: wf x0
H4: disjoint nil x0S (max_heap x0) > max_heap x0
free x
is nicely dual to new
.
p: ptr{{p -->?}} free p {{fun _ : unit => emp}}p: ptr{{p -->?}} free p {{fun _ : unit => emp}}p: ptr
x0: heap
x1: dynamic
H: wf ((p, x1) :: nil)
H0: wf x0
H4: disjoint ((p, x1) :: nil) x0match match lookup p (insert p x1 x0) with | Some _ => Some (remove p (insert p x1 x0), tt) | None => None end with | Some (h', _) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ h1 = nil /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endp: ptr
x0: heap
x1: dynamic
H: wf ((p, x1) :: nil)
H0: wf x0
H4: disjoint ((p, x1) :: nil) x0exists h1 h2 : heap, wf h1 /\ wf h2 /\ h1 = nil /\ x0 = h2 /\ remove p (insert p x1 x0) = merge h1 h2 /\ disjoint h1 h2p: ptr
x0: heap
x1: dynamic
H: wf ((p, x1) :: nil)
H0: wf x0
H4: disjoint ((p, x1) :: nil) x0exists h2 : heap, wf nil /\ wf h2 /\ nil = nil /\ x0 = h2 /\ remove p (insert p x1 x0) = merge nil h2 /\ disjoint nil h2p: ptr
x0: heap
x1: dynamic
H: wf ((p, x1) :: nil)
H0: wf x0
H4: disjoint ((p, x1) :: nil) x0wf nil /\ wf x0 /\ nil = nil /\ x0 = x0 /\ remove p (insert p x1 x0) = merge nil x0 /\ disjoint nil x0p: ptr
x0: heap
x1: dynamic
H: wf ((p, x1) :: nil)
H0: wf x0
H4: disjoint ((p, x1) :: nil) x0remove p (insert p x1 x0) = x0auto. Qed.p: ptr
x0: heap
x1: dynamic
H: wf ((p, x1) :: nil)
H0: wf x0
H4: disjoint ((p, x1) :: nil) x0x0 = x0
write p v
requires a heap where p
points to some value, and ensures p
points to v
afterwards.
t: stype
p: tptr t
v: interp t{{p -->?}} write p v {{fun _ : unit => p --> v}}t: stype
p: tptr t
v: interp t{{p -->?}} write p v {{fun _ : unit => p --> v}}t: stype
p: tptr t
v: interp tforall h h2 : heap, (exists h1 h3 : heap, wf h1 /\ wf h3 /\ (exists x : dynamic, h1 = (p, x) :: nil) /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match match lookup p h with | Some _ => Some (insert p (existT interp t v) (remove p h), tt) | None => None end with | Some (h', _) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ h1 = (p, existT interp t v) :: nil /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False endt: stype
p: tptr t
v: interp t
x0: heap
x1: dynamic
H: True /\ True
H0: wf x0
H4: ~ indom p x0 /\ Truematch match lookup p (insert p x1 x0) with | Some _ => Some (insert p (existT interp t v) (remove p (insert p x1 x0)), tt) | None => None end with | Some (h', _) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ h1 = (p, existT interp t v) :: nil /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endt: stype
p: tptr t
v: interp t
x0: heap
x1: dynamic
H: True /\ True
H0: wf x0
H4: ~ indom p x0 /\ Trueexists h1 h2 : heap, wf h1 /\ wf h2 /\ h1 = (p, existT interp t v) :: nil /\ x0 = h2 /\ insert p (existT interp t v) (remove p (insert p x1 x0)) = merge h1 h2 /\ disjoint h1 h2t: stype
p: tptr t
v: interp t
x0: heap
x1: dynamic
H: True /\ True
H0: wf x0
H4: ~ indom p x0 /\ Trueexists h2 : heap, wf ((p, existT interp t v) :: nil) /\ wf h2 /\ (p, existT interp t v) :: nil = (p, existT interp t v) :: nil /\ x0 = h2 /\ insert p (existT interp t v) (remove p (insert p x1 x0)) = merge ((p, existT interp t v) :: nil) h2 /\ disjoint ((p, existT interp t v) :: nil) h2t: stype
p: tptr t
v: interp t
x0: heap
x1: dynamic
H: True /\ True
H0: wf x0
H4: ~ indom p x0 /\ Truewf ((p, existT interp t v) :: nil) /\ wf x0 /\ (p, existT interp t v) :: nil = (p, existT interp t v) :: nil /\ x0 = x0 /\ insert p (existT interp t v) (remove p (insert p x1 x0)) = merge ((p, existT interp t v) :: nil) x0 /\ disjoint ((p, existT interp t v) :: nil) x0rewrite remove_insert ; auto. Qed.t: stype
p: tptr t
v: interp t
x0: heap
x1: dynamic
H, H3: True
H0: wf x0
H1: ~ indom p x0
H2: Trueinsert p (existT interp t v) (remove p (insert p x1 x0)) = insert p (existT interp t v) x0
read p
requires a heap where p
points to some value v
, and ensures p
points to v
afterwards, and the result is equal to v
.
t: stype
p: tptr t
v: interp t{{p --> v}} read p {{fun x : interp t => p --> x ** pure (x = v)}}t: stype
p: tptr t
v: interp t{{p --> v}} read p {{fun x : interp t => p --> x ** pure (x = v)}}t: stype
p: tptr t
v: interp t
x0: heap
H, H3: True
H0: wf x0
H1: ~ indom p x0
H2: Truematch match match lookup p (insert p (existT interp t v) x0) with | Some u => Some (insert p (existT interp t v) x0, u) | None => None end with | Some (h2, v) => (let (t', v0) := v in match stype_eq t t' with | left Heq => ret (eq_rec_r (fun t : stype => interp t) v0 Heq) | right _ => exit (interp t) end) h2 | None => None end with | Some (h', v0) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ h3 = (p, existT interp t v0) :: nil /\ (h4 = nil /\ v0 = v) /\ h1 = merge h3 h4 /\ disjoint h3 h4) /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endt: stype
p: tptr t
v: interp t
x0: heap
H, H3: True
H0: wf x0
H1: ~ indom p x0
H2: Truematch match stype_eq t t with | left Heq => ret (eq_rec_r (fun t : stype => interp t) v Heq) | right _ => exit (interp t) end (insert p (existT interp t v) x0) with | Some (h', v0) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ h3 = (p, existT interp t v0) :: nil /\ (h4 = nil /\ v0 = v) /\ h1 = merge h3 h4 /\ disjoint h3 h4) /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endt: stype
p: tptr t
v: interp t
x0: heap
H, H3: True
H0: wf x0
H1: ~ indom p x0
H2: True
e: t = tmatch ret (eq_rec_r (fun t : stype => interp t) v e) (insert p (existT interp t v) x0) with | Some (h', v0) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ h3 = (p, existT interp t v0) :: nil /\ (h4 = nil /\ v0 = v) /\ h1 = merge h3 h4 /\ disjoint h3 h4) /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endt: stype
p: tptr t
v: interp t
x0: heap
H, H3: True
H0: wf x0
H1: ~ indom p x0
H2: True
e: t = tmatch ret (eq_rec_r (fun t : stype => interp t) v eq_refl) (insert p (existT interp t v) x0) with | Some (h', v0) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ h3 = (p, existT interp t v0) :: nil /\ (h4 = nil /\ v0 = v) /\ h1 = merge h3 h4 /\ disjoint h3 h4) /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endt: stype
p: tptr t
v: interp t
x0: heap
H, H3: True
H0: wf x0
H1: ~ indom p x0
H2: True
e: t = texists h1 h2 : heap, wf h1 /\ wf h2 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ h3 = (p, existT interp t v) :: nil /\ (h4 = nil /\ v = v) /\ h1 = merge h3 h4 /\ disjoint h3 h4) /\ x0 = h2 /\ insert p (existT interp t v) x0 = merge h1 h2 /\ disjoint h1 h2t: stype
p: tptr t
v: interp t
x0: heap
H, H3: True
H0: wf x0
H1: ~ indom p x0
H2: True
e: t = texists h1 h2 : heap, wf h1 /\ wf h2 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ h3 = (p, existT interp t v) :: nil /\ (h4 = nil /\ v = v) /\ h1 = merge h3 h4 /\ disjoint h3 h4) /\ x0 = h2 /\ insert p (existT interp t v) x0 = merge h1 h2 /\ disjoint h1 h2t: stype
p: tptr t
v: interp t
x0: heap
H, H3: True
H0: wf x0
H1: ~ indom p x0
H2: True
e: t = texists h2 : heap, wf ((p, existT interp t v) :: nil) /\ wf h2 /\ (exists h1 h3 : heap, wf h1 /\ wf h3 /\ h1 = (p, existT interp t v) :: nil /\ (h3 = nil /\ v = v) /\ (p, existT interp t v) :: nil = merge h1 h3 /\ disjoint h1 h3) /\ x0 = h2 /\ insert p (existT interp t v) x0 = merge ((p, existT interp t v) :: nil) h2 /\ disjoint ((p, existT interp t v) :: nil) h2t: stype
p: tptr t
v: interp t
x0: heap
H, H3: True
H0: wf x0
H1: ~ indom p x0
H2: True
e: t = twf ((p, existT interp t v) :: nil) /\ wf x0 /\ (exists h1 h2 : heap, wf h1 /\ wf h2 /\ h1 = (p, existT interp t v) :: nil /\ (h2 = nil /\ v = v) /\ (p, existT interp t v) :: nil = merge h1 h2 /\ disjoint h1 h2) /\ x0 = x0 /\ insert p (existT interp t v) x0 = merge ((p, existT interp t v) :: nil) x0 /\ disjoint ((p, existT interp t v) :: nil) x0t: stype
p: tptr t
v: interp t
x0: heap
H, H3: True
H0: wf x0
H1: ~ indom p x0
H2: True
e: t = texists h1 h2 : heap, wf h1 /\ wf h2 /\ h1 = (p, existT interp t v) :: nil /\ (h2 = nil /\ v = v) /\ (p, existT interp t v) :: nil = merge h1 h2 /\ disjoint h1 h2t: stype
p: tptr t
v: interp t
x0: heap
H, H3: True
H0: wf x0
H1: ~ indom p x0
H2: True
e: t = texists h2 : heap, wf ((p, existT interp t v) :: nil) /\ wf h2 /\ (p, existT interp t v) :: nil = (p, existT interp t v) :: nil /\ (h2 = nil /\ v = v) /\ (p, existT interp t v) :: nil = merge ((p, existT interp t v) :: nil) h2 /\ disjoint ((p, existT interp t v) :: nil) h2mysimp. Qed.t: stype
p: tptr t
v: interp t
x0: heap
H, H3: True
H0: wf x0
H1: ~ indom p x0
H2: True
e: t = twf ((p, existT interp t v) :: nil) /\ wf nil /\ (p, existT interp t v) :: nil = (p, existT interp t v) :: nil /\ (nil = nil /\ v = v) /\ (p, existT interp t v) :: nil = merge ((p, existT interp t v) :: nil) nil /\ disjoint ((p, existT interp t v) :: nil) nil
This is one possible proof rule for bind
. Basically, we require that the post-condition of the first command implies the pre-condition of the second command.
T1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2{{P1}} c {{Q1}} -> (forall x : T1, {{P2 x}} f x {{Q2}}) -> (forall x : T1, Q1 x ==> P2 x) -> {{P1}} x <- c; f x {{Q2}}T1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2{{P1}} c {{Q1}} -> (forall x : T1, {{P2 x}} f x {{Q2}}) -> (forall x : T1, Q1 x ==> P2 x) -> {{P1}} x <- c; f x {{Q2}}T1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H: forall h h2 : heap, (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P1 h1 /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match c h with | Some (h', v) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q1 v h1 /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False end
H0: forall (x : T1) (h h2 : heap), (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P2 x h1 /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match f x h with | Some (h', v) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q2 v h1 /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False end
H1: forall x : T1, Q1 x ==> P2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P1 x
H7: disjoint x x0match match c (merge x x0) with | Some (h2, v) => f v h2 | None => None end with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H: forall h h2 : heap, (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P1 h1 /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match c h with | Some (h', v) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q1 v h1 /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False end
H0: forall (x : T1) (h h2 : heap), (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P2 x h1 /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match f x h with | Some (h', v) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q2 v h1 /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False end
H1: forall x : T1, Q1 x ==> P2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P1 x
H7: disjoint x x0((exists h1 h2 : heap, wf h1 /\ wf h2 /\ P1 h1 /\ x0 = h2 /\ merge x x0 = merge h1 h2 /\ disjoint h1 h2) -> match c (merge x x0) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q1 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end) -> match match c (merge x x0) with | Some (h2, v) => f v h2 | None => None end with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H0: forall (x : T1) (h h2 : heap), (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P2 x h1 /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match f x h with | Some (h', v) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q2 v h1 /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False end
H1: forall x : T1, Q1 x ==> P2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P1 x
H7: disjoint x x0((exists h1 h2 : heap, wf h1 /\ wf h2 /\ P1 h1 /\ x0 = h2 /\ merge x x0 = merge h1 h2 /\ disjoint h1 h2) -> match c (merge x x0) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q1 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end) -> match match c (merge x x0) with | Some (h2, v) => f v h2 | None => None end with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H0: forall (x : T1) (h h2 : heap), (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P2 x h1 /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match f x h with | Some (h', v) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q2 v h1 /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False end
H1: forall x : T1, Q1 x ==> P2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P1 x
H7: disjoint x x0exists x1 h2 : heap, wf x1 /\ wf h2 /\ P1 x1 /\ x0 = h2 /\ merge x x0 = merge x1 h2 /\ disjoint x1 h2T1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H0: forall (x : T1) (h h2 : heap), (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P2 x h1 /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match f x h with | Some (h', v) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q2 v h1 /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False end
H1: forall x : T1, Q1 x ==> P2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P1 x
H7: disjoint x x0match c (merge x x0) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q1 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end -> match match c (merge x x0) with | Some (h2, v) => f v h2 | None => None end with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H0: forall (x : T1) (h h2 : heap), (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P2 x h1 /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match f x h with | Some (h', v) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q2 v h1 /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False end
H1: forall x : T1, Q1 x ==> P2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P1 x
H7: disjoint x x0exists h2 : heap, wf x /\ wf h2 /\ P1 x /\ x0 = h2 /\ merge x x0 = merge x h2 /\ disjoint x h2T1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H0: forall (x : T1) (h h2 : heap), (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P2 x h1 /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match f x h with | Some (h', v) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q2 v h1 /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False end
H1: forall x : T1, Q1 x ==> P2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P1 x
H7: disjoint x x0match c (merge x x0) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q1 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end -> match match c (merge x x0) with | Some (h2, v) => f v h2 | None => None end with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H0: forall (x : T1) (h h2 : heap), (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P2 x h1 /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match f x h with | Some (h', v) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q2 v h1 /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False end
H1: forall x : T1, Q1 x ==> P2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P1 x
H7: disjoint x x0wf x /\ wf x0 /\ P1 x /\ x0 = x0 /\ merge x x0 = merge x x0 /\ disjoint x x0T1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H0: forall (x : T1) (h h2 : heap), (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P2 x h1 /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match f x h with | Some (h', v) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q2 v h1 /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False end
H1: forall x : T1, Q1 x ==> P2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P1 x
H7: disjoint x x0match c (merge x x0) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q1 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end -> match match c (merge x x0) with | Some (h2, v) => f v h2 | None => None end with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H0: forall (x : T1) (h h2 : heap), (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P2 x h1 /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match f x h with | Some (h', v) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q2 v h1 /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False end
H1: forall x : T1, Q1 x ==> P2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P1 x
H7: disjoint x x0match c (merge x x0) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q1 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end -> match match c (merge x x0) with | Some (h2, v) => f v h2 | None => None end with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H0: forall (x : T1) (h h2 : heap), (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P2 x h1 /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match f x h with | Some (h', v) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q2 v h1 /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False end
H1: forall x : T1, Q1 x ==> P2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P1 x
H7: disjoint x x0
h: heap
t: T1
x1, x2: heap
H: wf x1
H5: wf x2
H6: Q1 t x1
H8: x0 = x2
H9: h = merge x1 x2
H10: disjoint x1 x2match f t h with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H0: forall (x : T1) (h h2 : heap), (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P2 x h1 /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match f x h with | Some (h', v) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q2 v h1 /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False end
H1: forall x : T1, Q1 x ==> P2 x
x: heap
H2: wf x
x2: heap
H3: wf x2
H4: P1 x
H7: disjoint x x2
t: T1
x1: heap
H: wf x1
H5: wf x2
H6: Q1 t x1
H10: disjoint x1 x2match f t (merge x1 x2) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x2 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H0: forall (x : T1) (h h2 : heap), (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P2 x h1 /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match f x h with | Some (h', v) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q2 v h1 /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False end
H1: forall x : T1, Q1 x ==> P2 x
x: heap
H2: wf x
x2: heap
H3: wf x2
H4: P1 x
H7: disjoint x x2
t: T1
x1: heap
H: wf x1
H5: wf x2
H6: Q1 t x1
H10: disjoint x1 x2((exists h1 h2 : heap, wf h1 /\ wf h2 /\ P2 t h1 /\ x2 = h2 /\ merge x1 x2 = merge h1 h2 /\ disjoint h1 h2) -> match f t (merge x1 x2) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x2 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end) -> match f t (merge x1 x2) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x2 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H1: forall x : T1, Q1 x ==> P2 x
x: heap
H2: wf x
x2: heap
H3: wf x2
H4: P1 x
H7: disjoint x x2
t: T1
x1: heap
H: wf x1
H5: wf x2
H6: Q1 t x1
H10: disjoint x1 x2((exists h1 h2 : heap, wf h1 /\ wf h2 /\ P2 t h1 /\ x2 = h2 /\ merge x1 x2 = merge h1 h2 /\ disjoint h1 h2) -> match f t (merge x1 x2) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x2 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end) -> match f t (merge x1 x2) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x2 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H1: forall x : T1, Q1 x ==> P2 x
x: heap
H2: wf x
x2: heap
H3: wf x2
H4: P1 x
H7: disjoint x x2
t: T1
x1: heap
H: wf x1
H5: wf x2
H6: Q1 t x1
H10: disjoint x1 x2exists x h2 : heap, wf x /\ wf h2 /\ P2 t x /\ x2 = h2 /\ merge x1 x2 = merge x h2 /\ disjoint x h2T1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H1: forall x : T1, Q1 x ==> P2 x
x: heap
H2: wf x
x2: heap
H3: wf x2
H4: P1 x
H7: disjoint x x2
t: T1
x1: heap
H: wf x1
H5: wf x2
H6: Q1 t x1
H10: disjoint x1 x2match f t (merge x1 x2) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x2 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end -> match f t (merge x1 x2) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x2 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H1: forall x : T1, Q1 x ==> P2 x
x: heap
H2: wf x
x2: heap
H3: wf x2
H4: P1 x
H7: disjoint x x2
t: T1
x1: heap
H: wf x1
H5: wf x2
H6: Q1 t x1
H10: disjoint x1 x2exists h2 : heap, wf x1 /\ wf h2 /\ P2 t x1 /\ x2 = h2 /\ merge x1 x2 = merge x1 h2 /\ disjoint x1 h2T1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H1: forall x : T1, Q1 x ==> P2 x
x: heap
H2: wf x
x2: heap
H3: wf x2
H4: P1 x
H7: disjoint x x2
t: T1
x1: heap
H: wf x1
H5: wf x2
H6: Q1 t x1
H10: disjoint x1 x2match f t (merge x1 x2) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x2 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end -> match f t (merge x1 x2) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x2 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H1: forall x : T1, Q1 x ==> P2 x
x: heap
H2: wf x
x2: heap
H3: wf x2
H4: P1 x
H7: disjoint x x2
t: T1
x1: heap
H: wf x1
H5: wf x2
H6: Q1 t x1
H10: disjoint x1 x2wf x1 /\ wf x2 /\ P2 t x1 /\ x2 = x2 /\ merge x1 x2 = merge x1 x2 /\ disjoint x1 x2T1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H1: forall x : T1, Q1 x ==> P2 x
x: heap
H2: wf x
x2: heap
H3: wf x2
H4: P1 x
H7: disjoint x x2
t: T1
x1: heap
H: wf x1
H5: wf x2
H6: Q1 t x1
H10: disjoint x1 x2match f t (merge x1 x2) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x2 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end -> match f t (merge x1 x2) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x2 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H1: forall x : T1, Q1 x ==> P2 x
x: heap
H2: wf x
x2: heap
H3: wf x2
H4: P1 x
H7: disjoint x x2
t: T1
x1: heap
H: wf x1
H5: wf x2
H6: Q1 t x1
H10: disjoint x1 x2P2 t x1T1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H1: forall x : T1, Q1 x ==> P2 x
x: heap
H2: wf x
x2: heap
H3: wf x2
H4: P1 x
H7: disjoint x x2
t: T1
x1: heap
H: wf x1
H5: wf x2
H6: Q1 t x1
H10: disjoint x1 x2match f t (merge x1 x2) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x2 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end -> match f t (merge x1 x2) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x2 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endauto. Qed. Arguments bind_tc {T1 T2 P1 Q1 P2 Q2 c f}.T1, T2: Type
P1: hprop
Q1, P2: T1 -> hprop
Q2: T2 -> hprop
c: Cmd T1
f: T1 -> Cmd T2
H1: forall x : T1, Q1 x ==> P2 x
x: heap
H2: wf x
x2: heap
H3: wf x2
H4: P1 x
H7: disjoint x x2
t: T1
x1: heap
H: wf x1
H5: wf x2
H6: Q1 t x1
H10: disjoint x1 x2match f t (merge x1 x2) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x2 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end -> match f t (merge x1 x2) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x2 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end
And we also have a rule of consequence which allows us to strengthen the pre-condition and weaken the post-condition. Note however, that this will not allow us to "forget" any locations in our footprint.
T: Type
P1: hprop
Q1: T -> hprop
P2: hprop
Q2: T -> hprop
c: Cmd T{{P1}} c {{Q1}} -> P2 ==> P1 -> (forall x : T, Q1 x ==> Q2 x) -> {{P2}} c {{Q2}}T: Type
P1: hprop
Q1: T -> hprop
P2: hprop
Q2: T -> hprop
c: Cmd T{{P1}} c {{Q1}} -> P2 ==> P1 -> (forall x : T, Q1 x ==> Q2 x) -> {{P2}} c {{Q2}}T: Type
P1: hprop
Q1: T -> hprop
P2: hprop
Q2: T -> hprop
c: Cmd T
H: forall h h2 : heap, (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P1 h1 /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match c h with | Some (h', v) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q1 v h1 /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False end
H0: P2 ==> P1
H1: forall x : T, Q1 x ==> Q2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P2 x
H7: disjoint x x0match c (merge x x0) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT: Type
P1: hprop
Q1: T -> hprop
P2: hprop
Q2: T -> hprop
c: Cmd T
H: forall h h2 : heap, (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P1 h1 /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match c h with | Some (h', v) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q1 v h1 /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False end
H0: P2 ==> P1
H1: forall x : T, Q1 x ==> Q2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P2 x
H7: disjoint x x0((exists h1 h2 : heap, wf h1 /\ wf h2 /\ P1 h1 /\ x0 = h2 /\ merge x x0 = merge h1 h2 /\ disjoint h1 h2) -> match c (merge x x0) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q1 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end) -> match c (merge x x0) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT: Type
P1: hprop
Q1: T -> hprop
P2: hprop
Q2: T -> hprop
c: Cmd T
H0: P2 ==> P1
H1: forall x : T, Q1 x ==> Q2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P2 x
H7: disjoint x x0((exists h1 h2 : heap, wf h1 /\ wf h2 /\ P1 h1 /\ x0 = h2 /\ merge x x0 = merge h1 h2 /\ disjoint h1 h2) -> match c (merge x x0) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q1 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end) -> match c (merge x x0) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT: Type
P1: hprop
Q1: T -> hprop
P2: hprop
Q2: T -> hprop
c: Cmd T
H0: P2 ==> P1
H1: forall x : T, Q1 x ==> Q2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P2 x
H7: disjoint x x0exists x1 h2 : heap, wf x1 /\ wf h2 /\ P1 x1 /\ x0 = h2 /\ merge x x0 = merge x1 h2 /\ disjoint x1 h2T: Type
P1: hprop
Q1: T -> hprop
P2: hprop
Q2: T -> hprop
c: Cmd T
H0: P2 ==> P1
H1: forall x : T, Q1 x ==> Q2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P2 x
H7: disjoint x x0match c (merge x x0) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q1 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end -> match c (merge x x0) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT: Type
P1: hprop
Q1: T -> hprop
P2: hprop
Q2: T -> hprop
c: Cmd T
H0: P2 ==> P1
H1: forall x : T, Q1 x ==> Q2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P2 x
H7: disjoint x x0exists h2 : heap, wf x /\ wf h2 /\ P1 x /\ x0 = h2 /\ merge x x0 = merge x h2 /\ disjoint x h2T: Type
P1: hprop
Q1: T -> hprop
P2: hprop
Q2: T -> hprop
c: Cmd T
H0: P2 ==> P1
H1: forall x : T, Q1 x ==> Q2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P2 x
H7: disjoint x x0match c (merge x x0) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q1 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end -> match c (merge x x0) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT: Type
P1: hprop
Q1: T -> hprop
P2: hprop
Q2: T -> hprop
c: Cmd T
H0: P2 ==> P1
H1: forall x : T, Q1 x ==> Q2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P2 x
H7: disjoint x x0wf x /\ wf x0 /\ P1 x /\ x0 = x0 /\ merge x x0 = merge x x0 /\ disjoint x x0T: Type
P1: hprop
Q1: T -> hprop
P2: hprop
Q2: T -> hprop
c: Cmd T
H0: P2 ==> P1
H1: forall x : T, Q1 x ==> Q2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P2 x
H7: disjoint x x0match c (merge x x0) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q1 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end -> match c (merge x x0) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT: Type
P1: hprop
Q1: T -> hprop
P2: hprop
Q2: T -> hprop
c: Cmd T
H0: P2 ==> P1
H1: forall x : T, Q1 x ==> Q2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P2 x
H7: disjoint x x0match c (merge x x0) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q1 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end -> match c (merge x x0) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 v h1 /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT: Type
P1: hprop
Q1: T -> hprop
P2: hprop
Q2: T -> hprop
c: Cmd T
H0: P2 ==> P1
H1: forall x : T, Q1 x ==> Q2 x
x, x0: heap
H2: wf x
H3: wf x0
H4: P2 x
H7: disjoint x x0
h: heap
t: T
x1, x2: heap
H: wf x1
H5: wf x2
H6: Q1 t x1
H8: x0 = x2
H9: h = merge x1 x2
H10: disjoint x1 x2exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 t h1 /\ x0 = h2 /\ h = merge h1 h2 /\ disjoint h1 h2T: Type
P1: hprop
Q1: T -> hprop
P2: hprop
Q2: T -> hprop
c: Cmd T
H0: P2 ==> P1
H1: forall x : T, Q1 x ==> Q2 x
x: heap
H2: wf x
x2: heap
H3: wf x2
H4: P2 x
H7: disjoint x x2
t: T
x1: heap
H: wf x1
H5: wf x2
H6: Q1 t x1
H10: disjoint x1 x2exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q2 t h1 /\ x2 = h2 /\ merge x1 x2 = merge h1 h2 /\ disjoint h1 h2T: Type
P1: hprop
Q1: T -> hprop
P2: hprop
Q2: T -> hprop
c: Cmd T
H0: P2 ==> P1
H1: forall x : T, Q1 x ==> Q2 x
x: heap
H2: wf x
x2: heap
H3: wf x2
H4: P2 x
H7: disjoint x x2
t: T
x1: heap
H: wf x1
H5: wf x2
H6: Q1 t x1
H10: disjoint x1 x2exists h2 : heap, wf x1 /\ wf h2 /\ Q2 t x1 /\ x2 = h2 /\ merge x1 x2 = merge x1 h2 /\ disjoint x1 h2T: Type
P1: hprop
Q1: T -> hprop
P2: hprop
Q2: T -> hprop
c: Cmd T
H0: P2 ==> P1
H1: forall x : T, Q1 x ==> Q2 x
x: heap
H2: wf x
x2: heap
H3: wf x2
H4: P2 x
H7: disjoint x x2
t: T
x1: heap
H: wf x1
H5: wf x2
H6: Q1 t x1
H10: disjoint x1 x2wf x1 /\ wf x2 /\ Q2 t x1 /\ x2 = x2 /\ merge x1 x2 = merge x1 x2 /\ disjoint x1 x2eapply H1 ; mysimp. Qed. Arguments consequence_tc {T} _ _ {P2 Q2 c}.T: Type
P1: hprop
Q1: T -> hprop
P2: hprop
Q2: T -> hprop
c: Cmd T
H0: P2 ==> P1
H1: forall x : T, Q1 x ==> Q2 x
x: heap
H2: wf x
x2: heap
H3: wf x2
H4: P2 x
H7: disjoint x x2
t: T
x1: heap
H: wf x1
H5: wf x2
H6: Q1 t x1
H10: disjoint x1 x2Q2 t x1
This is a specialization of consequence that is a little more useful.
T: Type
P1, P2: hprop
Q1: T -> hprop
c: Cmd T{{P1}} c {{Q1}} -> P2 ==> P1 -> {{P2}} c {{Q1}}T: Type
P1, P2: hprop
Q1: T -> hprop
c: Cmd T{{P1}} c {{Q1}} -> P2 ==> P1 -> {{P2}} c {{Q1}}eapply consequence_tc ; eauto with sep_db. Qed. Arguments strengthen_tc {T} _ {P2 Q1 c}.T: Type
P1, P2: hprop
Q1: T -> hprop
c: Cmd T
H: {{P1}} c {{Q1}}
H0: P2 ==> P1{{P2}} c {{Q1}}
This is a specialization of consequence that is a little more useful.
T: Type
P1: hprop
Q1, Q2: T -> hprop
c: Cmd T{{P1}} c {{Q1}} -> (forall x : T, Q1 x ==> Q2 x) -> {{P1}} c {{Q2}}T: Type
P1: hprop
Q1, Q2: T -> hprop
c: Cmd T{{P1}} c {{Q1}} -> (forall x : T, Q1 x ==> Q2 x) -> {{P1}} c {{Q2}}eapply consequence_tc ; eauto with sep_db. Qed.T: Type
P1: hprop
Q1, Q2: T -> hprop
c: Cmd T
H: {{P1}} c {{Q1}}
H0: forall x : T, Q1 x ==> Q2 x{{P1}} c {{Q2}}
Finally, this is the most important rule and the one we lacked with Hoare logic: If {{ P }} c {{ Q }}
, then also {{ P ** R }} c {{ Q ** R }}
. That is, properties such as R
which are disjoint from the footprint are preserved when we run the command.
T: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop{{P}} c {{Q}} -> {{P ** R}} c {{fun x : T => Q x ** R}}T: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop{{P}} c {{Q}} -> {{P ** R}} c {{fun x : T => Q x ** R}}T: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
H: forall h h2 : heap, (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P h1 /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match c h with | Some (h', v) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q v h1 /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False end
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H5: disjoint (merge x1 x2) x0match c (merge (merge x1 x2) x0) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ Q v h3 /\ R h4 /\ h1 = merge h3 h4 /\ disjoint h3 h4) /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
H: forall h h2 : heap, (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P h1 /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match c h with | Some (h', v) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q v h1 /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False end
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H3: disjoint x0 x1
H4: disjoint x0 x2match c (merge x1 (merge x2 x0)) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ Q v h3 /\ R h4 /\ h1 = merge h3 h4 /\ disjoint h3 h4) /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
H: forall h h2 : heap, (exists h1 h3 : heap, wf h1 /\ wf h3 /\ P h1 /\ h2 = h3 /\ h = merge h1 h3 /\ disjoint h1 h3) -> match c h with | Some (h', v) => exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q v h1 /\ h2 = h3 /\ h' = merge h1 h3 /\ disjoint h1 h3 | None => False end
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H3: disjoint x0 x1
H4: disjoint x0 x2((exists h1 h2 : heap, wf h1 /\ wf h2 /\ P h1 /\ merge x2 x0 = h2 /\ merge x1 (merge x2 x0) = merge h1 h2 /\ disjoint h1 h2) -> match c (merge x1 (merge x2 x0)) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q v h1 /\ merge x2 x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end) -> match c (merge x1 (merge x2 x0)) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ Q v h3 /\ R h4 /\ h1 = merge h3 h4 /\ disjoint h3 h4) /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H3: disjoint x0 x1
H4: disjoint x0 x2((exists h1 h2 : heap, wf h1 /\ wf h2 /\ P h1 /\ merge x2 x0 = h2 /\ merge x1 (merge x2 x0) = merge h1 h2 /\ disjoint h1 h2) -> match c (merge x1 (merge x2 x0)) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q v h1 /\ merge x2 x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end) -> match c (merge x1 (merge x2 x0)) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ Q v h3 /\ R h4 /\ h1 = merge h3 h4 /\ disjoint h3 h4) /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H3: disjoint x0 x1
H4: disjoint x0 x2exists x h2 : heap, wf x /\ wf h2 /\ P x /\ merge x2 x0 = h2 /\ merge x1 (merge x2 x0) = merge x h2 /\ disjoint x h2T: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H3: disjoint x0 x1
H4: disjoint x0 x2match c (merge x1 (merge x2 x0)) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q v h1 /\ merge x2 x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end -> match c (merge x1 (merge x2 x0)) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ Q v h3 /\ R h4 /\ h1 = merge h3 h4 /\ disjoint h3 h4) /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H3: disjoint x0 x1
H4: disjoint x0 x2exists h2 : heap, wf x1 /\ wf h2 /\ P x1 /\ merge x2 x0 = h2 /\ merge x1 (merge x2 x0) = merge x1 h2 /\ disjoint x1 h2T: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H3: disjoint x0 x1
H4: disjoint x0 x2match c (merge x1 (merge x2 x0)) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q v h1 /\ merge x2 x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end -> match c (merge x1 (merge x2 x0)) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ Q v h3 /\ R h4 /\ h1 = merge h3 h4 /\ disjoint h3 h4) /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H3: disjoint x0 x1
H4: disjoint x0 x2wf x1 /\ wf (merge x2 x0) /\ P x1 /\ merge x2 x0 = merge x2 x0 /\ merge x1 (merge x2 x0) = merge x1 (merge x2 x0) /\ disjoint x1 (merge x2 x0)T: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H3: disjoint x0 x1
H4: disjoint x0 x2match c (merge x1 (merge x2 x0)) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q v h1 /\ merge x2 x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end -> match c (merge x1 (merge x2 x0)) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ Q v h3 /\ R h4 /\ h1 = merge h3 h4 /\ disjoint h3 h4) /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H3: disjoint x0 x1
H4: disjoint x0 x2match c (merge x1 (merge x2 x0)) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q v h1 /\ merge x2 x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False end -> match c (merge x1 (merge x2 x0)) with | Some (h', v) => exists h1 h2 : heap, wf h1 /\ wf h2 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ Q v h3 /\ R h4 /\ h1 = merge h3 h4 /\ disjoint h3 h4) /\ x0 = h2 /\ h' = merge h1 h2 /\ disjoint h1 h2 | None => False endT: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H3: disjoint x0 x1
H4: disjoint x0 x2
t: T
x: heap
H: wf x
H5: wf (merge x2 x0)
H9: Q t x
H13: disjoint x (merge x2 x0)exists h1 h2 : heap, wf h1 /\ wf h2 /\ (exists h3 h4 : heap, wf h3 /\ wf h4 /\ Q t h3 /\ R h4 /\ h1 = merge h3 h4 /\ disjoint h3 h4) /\ x0 = h2 /\ merge x (merge x2 x0) = merge h1 h2 /\ disjoint h1 h2T: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H3: disjoint x0 x1
H4: disjoint x0 x2
t: T
x: heap
H: wf x
H5: wf (merge x2 x0)
H9: Q t x
H13: disjoint x (merge x2 x0)exists h2 : heap, wf (merge x x2) /\ wf h2 /\ (exists h1 h3 : heap, wf h1 /\ wf h3 /\ Q t h1 /\ R h3 /\ merge x x2 = merge h1 h3 /\ disjoint h1 h3) /\ x0 = h2 /\ merge x (merge x2 x0) = merge (merge x x2) h2 /\ disjoint (merge x x2) h2T: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H3: disjoint x0 x1
H4: disjoint x0 x2
t: T
x: heap
H: wf x
H5: wf (merge x2 x0)
H9: Q t x
H13: disjoint x (merge x2 x0)wf (merge x x2) /\ wf x0 /\ (exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q t h1 /\ R h2 /\ merge x x2 = merge h1 h2 /\ disjoint h1 h2) /\ x0 = x0 /\ merge x (merge x2 x0) = merge (merge x x2) x0 /\ disjoint (merge x x2) x0T: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H3: disjoint x0 x1
H4: disjoint x0 x2
t: T
x: heap
H: wf x
H5: wf (merge x2 x0)
H9: Q t x
H11: disjoint x x2
H12: disjoint x x0exists h1 h2 : heap, wf h1 /\ wf h2 /\ Q t h1 /\ R h2 /\ merge x x2 = merge h1 h2 /\ disjoint h1 h2T: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H3: disjoint x0 x1
H4: disjoint x0 x2
t: T
x: heap
H: wf x
H5: wf (merge x2 x0)
H9: Q t x
H11: disjoint x x2
H12: disjoint x x0merge x (merge x2 x0) = merge (merge x x2) x0T: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H3: disjoint x0 x1
H4: disjoint x0 x2
t: T
x: heap
H: wf x
H5: wf (merge x2 x0)
H9: Q t x
H11: disjoint x x2
H12: disjoint x x0exists h2 : heap, wf x /\ wf h2 /\ Q t x /\ R h2 /\ merge x x2 = merge x h2 /\ disjoint x h2T: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H3: disjoint x0 x1
H4: disjoint x0 x2
t: T
x: heap
H: wf x
H5: wf (merge x2 x0)
H9: Q t x
H11: disjoint x x2
H12: disjoint x x0merge x (merge x2 x0) = merge (merge x x2) x0T: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H3: disjoint x0 x1
H4: disjoint x0 x2
t: T
x: heap
H: wf x
H5: wf (merge x2 x0)
H9: Q t x
H11: disjoint x x2
H12: disjoint x x0wf x /\ wf x2 /\ Q t x /\ R x2 /\ merge x x2 = merge x x2 /\ disjoint x x2T: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H3: disjoint x0 x1
H4: disjoint x0 x2
t: T
x: heap
H: wf x
H5: wf (merge x2 x0)
H9: Q t x
H11: disjoint x x2
H12: disjoint x x0merge x (merge x2 x0) = merge (merge x x2) x0rewrite merge_assoc ; disj ; auto. Qed. Arguments frame_tc {T c P Q}. Ltac sep := repeat (simpl in * ; (try subst) ; match goal with | [ |- forall _, _ ] => intros | [ |- _ ** ?Q ==> _ ** ?Q ] => eapply himp_split | [ |- ?P ** _ ==> ?P ** _ ] => eapply himp_split | [ |- ?P ** _ ==> _ ** ?P ] => eapply hyp_comm | [ |- hexists _ ==> _ ] => eapply hyp_hexists | [ |- _ ** pure _ ==> _ ] => eapply hyp_comm ; eapply hyp_pure | [ |- (_ ** _) ** _ ==> _ ] => eapply hyp_assoc_r | [ |- _ ==> (_ ** _) ** _ ] => eapply assoc_concl | _ => eauto with sep_db end).T: Type
c: Cmd T
P: hprop
Q: T -> hprop
R: hprop
x0, x1, x2: heap
H0: wf (merge x1 x2)
H1: wf x0
H2: wf x1
H6: wf x2
H7: P x1
H8: R x2
H10: disjoint x1 x2
H3: disjoint x0 x1
H4: disjoint x0 x2
t: T
x: heap
H: wf x
H5: wf (merge x2 x0)
H9: Q t x
H11: disjoint x x2
H12: disjoint x x0merge x (merge x2 x0) = merge (merge x x2) x0
Following are a few examples that illustrate the use of Separation Logic.
Definition inc(p:tptr Nat_t) := v <- read p ; write p (1 + v).
The next definition says that if we start in a state where p
holds n
,
then after running inc, we do not fail and get into a state
where p
holds 1+n
. Notice that it's rather delicate to
hang on to the fact that the value read out is equal to n
. If
we used binary post-conditions (a relation on both the input heap
and output heap), this wouldn't be necessary.
p: tptr Nat_t
n: interp Nat_t{{p --> n}} inc p {{fun _ : unit => p --> 1 + n}}p: tptr Nat_t
n: interp Nat_t{{p --> n}} inc p {{fun _ : unit => p --> 1 + n}}p: tptr Nat_t
n: nat{{p --> n}} v <- read p; write p (S v) {{fun _ : unit => p --> S n}}p: tptr Nat_t
n: nat{{p --> n}} read p {{?Q1}}p: tptr Nat_t
n, x: nat{{?Q1 x}} write p (S x) {{fun _ : unit => p --> S n}}p: tptr Nat_t
n, x: nat{{(fun x : interp Nat_t => p --> x ** pure (x = n)) x}} write p (S x) {{fun _ : unit => p --> S n}}p: tptr Nat_t
n, x: nat{{p --> x ** pure (x = n)}} write p (S x) {{fun _ : unit => p --> S n}}p: tptr Nat_t
n, x: nat{{?P1}} write p (S x) {{?Q1}}p: tptr Nat_t
n, x: natp --> x ** pure (x = n) ==> ?P1p: tptr Nat_t
n, x: natforall x0 : unit, ?Q1 x0 ==> p --> S np: tptr Nat_t
n, x: nat{{?P}} write p (S x) {{?Q}}p: tptr Nat_t
n, x: natp --> x ** pure (x = n) ==> ?P ** pure (x = n)p: tptr Nat_t
n, x: natforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = n)) x0 ==> p --> S np: tptr Nat_t
n, x: natp --> x ** pure (x = n) ==> p -->? ** pure (x = n)p: tptr Nat_t
n, x: natforall x0 : unit, (fun x1 : unit => (fun _ : unit => p --> S x) x1 ** pure (x = n)) x0 ==> p --> S np: tptr Nat_t
n, x: natforall x0 : unit, (fun x1 : unit => (fun _ : unit => p --> S x) x1 ** pure (x = n)) x0 ==> p --> S nsep. Qed.p: tptr Nat_t
n, x: natunit -> p --> S x ** pure (x = n) ==> p --> S n
The great part is that now if we have a property on some disjoint
part of the state, say p2 --> n2
, then after calling inc
, we
are guaranteed that property is preserved via the frame rule.
p1, p2: tptr Nat_t
n1, n2: interp Nat_t{{p1 --> n1 ** p2 --> n2}} inc p1 {{fun _ : unit => p1 --> 1 + n1 ** p2 --> n2}}p1, p2: tptr Nat_t
n1, n2: interp Nat_t{{p1 --> n1 ** p2 --> n2}} inc p1 {{fun _ : unit => p1 --> 1 + n1 ** p2 --> n2}}p1, p2: tptr Nat_t
n1, n2: interp Nat_t{{p1 --> n1 ** p2 --> n2}} inc p1 {{fun _ : unit => p1 --> 1 + n1 ** p2 --> n2}}apply inc_tc. Qed.p1, p2: tptr Nat_t
n1, n2: interp Nat_t{{p1 --> n1}} inc p1 {{fun _ : unit => p1 --> 1 + n1}}
Here is a function to swap the contents of two pointers.
Definition swap(t:stype)(p1 p2:tptr t) := v1 <- read p1 ; v2 <- read p2 ; write p2 v1 ;; write p1 v2.
Alas, reasoning isn't quite as simple as we might hope. We have to not only put in the right uses of the frame and consequence rules, but we must also so a lot of commuting and re-associating to discharge the verification conditions. For Ynot, Adam Chlipala wrote an Ltac tactic that mostly took care of this sort of simple reasoning. And Gonthier et al. have done some nice work showing how to use type-classes or canonical-structures to automate a lot of this sort of thing. Below, we'll see you an alternative technique based on reflection.
t: stype
p1, p2: tptr t
v1, v2: interp t{{p1 --> v1 ** p2 --> v2}} swap p1 p2 {{fun _ : unit => p1 --> v2 ** p2 --> v1}}t: stype
p1, p2: tptr t
v1, v2: interp t{{p1 --> v1 ** p2 --> v2}} swap p1 p2 {{fun _ : unit => p1 --> v2 ** p2 --> v1}}t: stype
p1, p2: tptr t
v1, v2: interp t{{p1 --> v1 ** p2 --> v2}} v1 <- read p1; v2 <- read p2; write p2 v1;; write p1 v2 {{fun _ : unit => p1 --> v2 ** p2 --> v1}}t: stype
p1, p2: tptr t
v1, v2: interp t{{p1 --> v1 ** p2 --> v2}} read p1 {{?Q1}}t: stype
p1, p2: tptr t
v1, v2, x: interp t{{?Q1 x}} v2 <- read p2; write p2 x;; write p1 v2 {{fun _ : unit => p1 --> v2 ** p2 --> v1}}t: stype
p1, p2: tptr t
v1, v2: interp t{{p1 --> v1}} read p1 {{?Q}}t: stype
p1, p2: tptr t
v1, v2, x: interp t{{(fun x : interp t => ?Q x ** p2 --> v2) x}} v2 <- read p2; write p2 x;; write p1 v2 {{fun _ : unit => p1 --> v2 ** p2 --> v1}}t: stype
p1, p2: tptr t
v1, v2, x: interp t{{(fun x : interp t => (fun x0 : interp t => p1 --> x0 ** pure (x0 = v1)) x ** p2 --> v2) x}} v2 <- read p2; write p2 x;; write p1 v2 {{fun _ : unit => p1 --> v2 ** p2 --> v1}}t: stype
p1, p2: tptr t
v1, v2, x: interp t{{(p1 --> x ** pure (x = v1)) ** p2 --> v2}} v2 <- read p2; write p2 x;; write p1 v2 {{fun _ : unit => p1 --> v2 ** p2 --> v1}}t: stype
p1, p2: tptr t
v1, v2, x: interp t{{(p2 --> v2 ** p1 --> x) ** pure (x = v1)}} v2 <- read p2; write p2 x;; write p1 v2 {{?Q1}}t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, ?Q1 x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x: interp t{{p2 --> v2 ** p1 --> x}} v2 <- read p2; write p2 x;; write p1 v2 {{?Q}}t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x: interp t{{p2 --> v2 ** p1 --> x}} read p2 {{?Q1}}t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t{{?Q1 x0}} write p2 x;; write p1 x0 {{?Q}}t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x: interp t{{p2 --> v2}} read p2 {{?Q0}}t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t{{(fun x0 : interp t => ?Q0 x0 ** p1 --> x) x0}} write p2 x;; write p1 x0 {{?Q}}t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t{{(fun x0 : interp t => (fun x : interp t => p2 --> x ** pure (x = v2)) x0 ** p1 --> x) x0}} write p2 x;; write p1 x0 {{?Q}}t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t{{(p2 --> x0 ** pure (x0 = v2)) ** p1 --> x}} write p2 x;; write p1 x0 {{?Q}}t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t{{(p2 -->? ** p1 --> x) ** pure (x0 = v2)}} write p2 x;; write p1 x0 {{?Q1}}t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t(p2 --> x0 ** pure (x0 = v2)) ** p1 --> x ==> (p2 -->? ** p1 --> x) ** pure (x0 = v2)t: stype
p1, p2: tptr t
v1, v2, x, x0: interp tforall x1 : unit, ?Q1 x1 ==> ?Q x1t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t{{p2 -->? ** p1 --> x}} write p2 x;; write p1 x0 {{?Q0}}t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t(p2 --> x0 ** pure (x0 = v2)) ** p1 --> x ==> (p2 -->? ** p1 --> x) ** pure (x0 = v2)t: stype
p1, p2: tptr t
v1, v2, x, x0: interp tforall x1 : unit, (fun x2 : unit => ?Q0 x2 ** pure (x0 = v2)) x1 ==> ?Q x1t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t{{p2 -->? ** p1 --> x}} write p2 x {{?Q1}}t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t
x1: unit{{?Q1 x1}} write p1 x0 {{?Q0}}t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t(p2 --> x0 ** pure (x0 = v2)) ** p1 --> x ==> (p2 -->? ** p1 --> x) ** pure (x0 = v2)t: stype
p1, p2: tptr t
v1, v2, x, x0: interp tforall x1 : unit, (fun x2 : unit => ?Q0 x2 ** pure (x0 = v2)) x1 ==> ?Q x1t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t{{p2 -->?}} write p2 x {{?Q1}}t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t
x1: unit{{(fun x1 : unit => ?Q1 x1 ** p1 --> x) x1}} write p1 x0 {{?Q0}}t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t(p2 --> x0 ** pure (x0 = v2)) ** p1 --> x ==> (p2 -->? ** p1 --> x) ** pure (x0 = v2)t: stype
p1, p2: tptr t
v1, v2, x, x0: interp tforall x1 : unit, (fun x2 : unit => ?Q0 x2 ** pure (x0 = v2)) x1 ==> ?Q x1t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t
x1: unit{{(fun x0 : unit => (fun _ : unit => p2 --> x) x0 ** p1 --> x) x1}} write p1 x0 {{?Q0}}t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t(p2 --> x0 ** pure (x0 = v2)) ** p1 --> x ==> (p2 -->? ** p1 --> x) ** pure (x0 = v2)t: stype
p1, p2: tptr t
v1, v2, x, x0: interp tforall x1 : unit, (fun x2 : unit => ?Q0 x2 ** pure (x0 = v2)) x1 ==> ?Q x1t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t
x1: unit{{p2 --> x ** p1 --> x}} write p1 x0 {{?Q0}}t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t(p2 --> x0 ** pure (x0 = v2)) ** p1 --> x ==> (p2 -->? ** p1 --> x) ** pure (x0 = v2)t: stype
p1, p2: tptr t
v1, v2, x, x0: interp tforall x1 : unit, (fun x2 : unit => ?Q0 x2 ** pure (x0 = v2)) x1 ==> ?Q x1t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t
x1: unit{{p1 -->? ** p2 --> x}} write p1 x0 {{?Q1}}t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t
x1: unitp2 --> x ** p1 --> x ==> p1 -->? ** p2 --> xt: stype
p1, p2: tptr t
v1, v2, x, x0: interp t
x1: unitforall x2 : unit, ?Q1 x2 ==> ?Q0 x2t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t(p2 --> x0 ** pure (x0 = v2)) ** p1 --> x ==> (p2 -->? ** p1 --> x) ** pure (x0 = v2)t: stype
p1, p2: tptr t
v1, v2, x, x0: interp tforall x1 : unit, (fun x2 : unit => ?Q0 x2 ** pure (x0 = v2)) x1 ==> ?Q x1t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t
x1: unit{{p1 -->?}} write p1 x0 {{?Q1}}t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t
x1: unitp2 --> x ** p1 --> x ==> p1 -->? ** p2 --> xt: stype
p1, p2: tptr t
v1, v2, x, x0: interp t
x1: unitforall x2 : unit, (fun x3 : unit => ?Q1 x3 ** p2 --> x) x2 ==> ?Q0 x2t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t(p2 --> x0 ** pure (x0 = v2)) ** p1 --> x ==> (p2 -->? ** p1 --> x) ** pure (x0 = v2)t: stype
p1, p2: tptr t
v1, v2, x, x0: interp tforall x1 : unit, (fun x2 : unit => ?Q0 x2 ** pure (x0 = v2)) x1 ==> ?Q x1t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t
x1: unitp2 --> x ** p1 --> x ==> p1 -->? ** p2 --> xt: stype
p1, p2: tptr t
v1, v2, x, x0: interp t
x1: unitforall x1 : unit, (fun x2 : unit => (fun _ : unit => p1 --> x0) x2 ** p2 --> x) x1 ==> ?Q0 x1t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t(p2 --> x0 ** pure (x0 = v2)) ** p1 --> x ==> (p2 -->? ** p1 --> x) ** pure (x0 = v2)t: stype
p1, p2: tptr t
v1, v2, x, x0: interp tforall x1 : unit, (fun x2 : unit => ?Q0 x2 ** pure (x0 = v2)) x1 ==> ?Q x1t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t
x1: unitp1 --> x ** p2 --> x ==> p1 -->? ** p2 --> xt: stype
p1, p2: tptr t
v1, v2, x, x0: interp t
x1: unitforall x1 : unit, (fun x2 : unit => (fun _ : unit => p1 --> x0) x2 ** p2 --> x) x1 ==> ?Q0 x1t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t(p2 --> x0 ** pure (x0 = v2)) ** p1 --> x ==> (p2 -->? ** p1 --> x) ** pure (x0 = v2)t: stype
p1, p2: tptr t
v1, v2, x, x0: interp tforall x1 : unit, (fun x2 : unit => ?Q0 x2 ** pure (x0 = v2)) x1 ==> ?Q x1t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t
x1: unitforall x1 : unit, (fun x2 : unit => (fun _ : unit => p1 --> x0) x2 ** p2 --> x) x1 ==> ?Q0 x1t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t(p2 --> x0 ** pure (x0 = v2)) ** p1 --> x ==> (p2 -->? ** p1 --> x) ** pure (x0 = v2)t: stype
p1, p2: tptr t
v1, v2, x, x0: interp tforall x1 : unit, (fun x2 : unit => ?Q0 x2 ** pure (x0 = v2)) x1 ==> ?Q x1t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t
x1, x2: unitp1 --> x0 ** p2 --> x ==> ?Q0 x2t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t(p2 --> x0 ** pure (x0 = v2)) ** p1 --> x ==> (p2 -->? ** p1 --> x) ** pure (x0 = v2)t: stype
p1, p2: tptr t
v1, v2, x, x0: interp tforall x1 : unit, (fun x2 : unit => ?Q0 x2 ** pure (x0 = v2)) x1 ==> ?Q x1t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x, x0: interp t(p2 --> x0 ** pure (x0 = v2)) ** p1 --> x ==> (p2 -->? ** p1 --> x) ** pure (x0 = v2)t: stype
p1, p2: tptr t
v1, v2, x, x0: interp tforall x1 : unit, (fun x2 : unit => (fun _ : unit => p1 --> x0 ** p2 --> x) x2 ** pure (x0 = v2)) x1 ==> ?Q x1t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x, x0: interp tp2 --> x0 ** pure (x0 = v2) ** p1 --> x ==> p2 -->? ** p1 --> x ** pure (x0 = v2)t: stype
p1, p2: tptr t
v1, v2, x, x0: interp tforall x1 : unit, (fun x2 : unit => (fun _ : unit => p1 --> x0 ** p2 --> x) x2 ** pure (x0 = v2)) x1 ==> ?Q x1t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x, x0: interp tforall x1 : unit, (fun x2 : unit => (fun _ : unit => p1 --> x0 ** p2 --> x) x2 ** pure (x0 = v2)) x1 ==> ?Q x1t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x: interp t
x1: unitp1 --> v2 ** p2 --> x ==> ?Q x1t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => ?Q x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x: interp t(p1 --> x ** pure (x = v1)) ** p2 --> v2 ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => (fun _ : unit => p1 --> v2 ** p2 --> x) x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1t: stype
p1, p2: tptr t
v1, v2, x: interp tp2 --> v2 ** p1 --> x ** pure (x = v1) ==> (p2 --> v2 ** p1 --> x) ** pure (x = v1)t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => (fun _ : unit => p1 --> v2 ** p2 --> x) x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1sep. Qed.t: stype
p1, p2: tptr t
v1, v2, x: interp tforall x0 : unit, (fun x1 : unit => (fun _ : unit => p1 --> v2 ** p2 --> x) x1 ** pure (x = v1)) x0 ==> p1 --> v2 ** p2 --> v1
Our goal will be to build a Coq function that can simplify a separation implication P ==> Q
and prove that function correct. One simple strategy is to flatten out all of the stars and cross off all of the simple predicates that appear in both P
and Q
.
But of course, we can't just crawl over an hprop
because in general, these are functions! So our trick is to write down some syntax that represents a particular predicate, and then give an interpretation that maps that syntax back to our real predicate. Then we can compute with the syntax.
The Quote library should help us with this, but alas, it's not as clever as we might hope it would be...
Definition hprop_name := nat. Definition hprop_map := list (hprop_name * hprop). Definition hprop_name_eq := PeanoNat.Nat.eq_dec.
We begin by giving a basic syntax for predicates. Atom
will be used to represent things like abstract predicates (e.g., a variable P
) or a points-to-predicate etc. We'll use an environment to map an hprop_name
to the real hprop
.
Inductive Hprop : Set := | Emp : Hprop | Atom : hprop_name -> Hprop | Star : Hprop -> Hprop -> Hprop. Infix "#" := Star (right associativity, at level 80) : sep_scope. Fixpoint lookup_hprop (n:hprop_name) (hp:hprop_map) := match hp with | nil => (pure False) | (m,P)::rest => if hprop_name_eq n m then P else lookup_hprop n rest end. Section HINTERP.
We'll assume we have an hprop_map
lying around.
Variable hmap : hprop_map.
Now we given an interpretation to the syntax for predicates:
Fixpoint hinterp (hp:Hprop) : hprop := match hp with | Emp => emp | Atom n => lookup_hprop n hmap | Star h1 h2 => star (hinterp h1) (hinterp h2) end.
We can flatten a predicate into a list of hprop_name
.
Fixpoint flatten (hp:Hprop) : list hprop_name := match hp with | Emp => nil | Atom n => n::nil | Star h1 h2 => (flatten h1) ++ (flatten h2) end.
This function removes exactly one copy of n
, an hprop_name
, from the input list, returning None
if we fail to find a copy of n
.
Fixpoint remove_one(n:hprop_name)(hps : list hprop_name) :option (list hprop_name) := match hps with | nil => None | (m::rest) => if hprop_name_eq n m then Some rest else match remove_one n rest with | None => None | Some hps' => Some (m::hps') end end.
This is the heart of our simplification algorithm. Here, we are running through the list of names hp1
, trying to cross off each one that occurs in hp2
. If the current name doesn't occur, we add it to the end of hp0
so that we can keep track of it. That is, our invariant at each step should be that hp0 ** hp1 ==> hp2
once we map the list back to an hprop
.
Fixpoint simplify (hp0 hp1 hp2:list hprop_name) := match hp1 with | nil => (hp0,hp2) | (n::hp1') => match remove_one n hp2 with | Some hp2' => simplify hp0 hp1' hp2' | None => simplify (hp0 ++ n::nil) hp1' hp2 end end.
Convert a list of names back into an hprop
.
Definition collapse := List.fold_right (fun n p => Star (Atom n) p) Emp.
Convert a list of names into an hprop
.
Definition interp_list := List.fold_right (fun n p => (lookup_hprop n hmap) ** p) emp.
Our cross-off algorithm takes two hprop
, flattens them into lists of names, simplifies the two lists by crossing off common names, and then returns the two hprop
we get by collapsing the resulting simplified lists.
Definition cross_off (hp1 hp2 : Hprop) : Hprop * Hprop := let (hp1', hp2') := simplify nil (flatten hp1) (flatten hp2) in (collapse hp1', collapse hp2').hmap: hprop_map
hp: list hprop_namehinterp (collapse hp) = interp_list hphmap: hprop_map
hp: list hprop_namehinterp (collapse hp) = interp_list hphmap: hprop_map
a: hprop_name
hp: list hprop_name
IHhp: hinterp (collapse hp) = interp_list hp(lookup_hprop a hmap ** hinterp (collapse hp)) = (lookup_hprop a hmap ** interp_list hp)auto. Qed.hmap: hprop_map
a: hprop_name
hp: list hprop_name
IHhp: hinterp (collapse hp) = interp_list hp(lookup_hprop a hmap ** interp_list hp) = (lookup_hprop a hmap ** interp_list hp)
The following are various lemmas needed to reason about the interpretation of the syntax.
hmap: hprop_map
hp1, hp2: list hprop_nameinterp_list (hp1 ++ hp2) ==> interp_list hp1 ** interp_list hp2hmap: hprop_map
hp1, hp2: list hprop_nameinterp_list (hp1 ++ hp2) ==> interp_list hp1 ** interp_list hp2hmap: hprop_map
hp2: list hprop_nameinterp_list (nil ++ hp2) ==> interp_list nil ** interp_list hp2hmap: hprop_map
a: hprop_name
hp1, hp2: list hprop_name
IHhp1: interp_list (hp1 ++ hp2) ==> interp_list hp1 ** interp_list hp2interp_list ((a :: hp1) ++ hp2) ==> interp_list (a :: hp1) ** interp_list hp2hmap: hprop_map
hp2: list hprop_nameinterp_list hp2 ==> emp ** interp_list hp2hmap: hprop_map
a: hprop_name
hp1, hp2: list hprop_name
IHhp1: interp_list (hp1 ++ hp2) ==> interp_list hp1 ** interp_list hp2interp_list ((a :: hp1) ++ hp2) ==> interp_list (a :: hp1) ** interp_list hp2hmap: hprop_map
hp2: list hprop_nameinterp_list hp2 ==> emp ** interp_list hp2hmap: hprop_map
a: hprop_name
hp1, hp2: list hprop_name
IHhp1: interp_list (hp1 ++ hp2) ==> interp_list hp1 ** interp_list hp2interp_list ((a :: hp1) ++ hp2) ==> interp_list (a :: hp1) ** interp_list hp2hmap: hprop_map
hp2: list hprop_nameinterp_list hp2 ==> interp_list hp2 ** emphmap: hprop_map
a: hprop_name
hp1, hp2: list hprop_name
IHhp1: interp_list (hp1 ++ hp2) ==> interp_list hp1 ** interp_list hp2interp_list ((a :: hp1) ++ hp2) ==> interp_list (a :: hp1) ** interp_list hp2hmap: hprop_map
hp2: list hprop_nameinterp_list hp2 ==> interp_list hp2hmap: hprop_map
a: hprop_name
hp1, hp2: list hprop_name
IHhp1: interp_list (hp1 ++ hp2) ==> interp_list hp1 ** interp_list hp2interp_list ((a :: hp1) ++ hp2) ==> interp_list (a :: hp1) ** interp_list hp2hmap: hprop_map
a: hprop_name
hp1, hp2: list hprop_name
IHhp1: interp_list (hp1 ++ hp2) ==> interp_list hp1 ** interp_list hp2interp_list ((a :: hp1) ++ hp2) ==> interp_list (a :: hp1) ** interp_list hp2hmap: hprop_map
a: hprop_name
hp1, hp2: list hprop_name
IHhp1: interp_list (hp1 ++ hp2) ==> interp_list hp1 ** interp_list hp2lookup_hprop a hmap ** interp_list (hp1 ++ hp2) ==> (lookup_hprop a hmap ** interp_list hp1) ** interp_list hp2sep. Qed.hmap: hprop_map
a: hprop_name
hp1, hp2: list hprop_name
IHhp1: interp_list (hp1 ++ hp2) ==> interp_list hp1 ** interp_list hp2lookup_hprop a hmap ** interp_list (hp1 ++ hp2) ==> (lookup_hprop a hmap ** interp_list hp1) ** interp_list hp2hmap: hprop_map
hp1, hp2: list hprop_nameinterp_list hp1 ** interp_list hp2 ==> interp_list (hp1 ++ hp2)hmap: hprop_map
hp1, hp2: list hprop_nameinterp_list hp1 ** interp_list hp2 ==> interp_list (hp1 ++ hp2)hmap: hprop_map
hp2: list hprop_nameinterp_list nil ** interp_list hp2 ==> interp_list (nil ++ hp2)hmap: hprop_map
a: hprop_name
hp1, hp2: list hprop_name
IHhp1: interp_list hp1 ** interp_list hp2 ==> interp_list (hp1 ++ hp2)interp_list (a :: hp1) ** interp_list hp2 ==> interp_list ((a :: hp1) ++ hp2)hmap: hprop_map
hp2: list hprop_nameemp ** interp_list hp2 ==> interp_list hp2hmap: hprop_map
a: hprop_name
hp1, hp2: list hprop_name
IHhp1: interp_list hp1 ** interp_list hp2 ==> interp_list (hp1 ++ hp2)interp_list (a :: hp1) ** interp_list hp2 ==> interp_list ((a :: hp1) ++ hp2)hmap: hprop_map
hp2: list hprop_nameemp ** interp_list hp2 ==> interp_list hp2hmap: hprop_map
a: hprop_name
hp1, hp2: list hprop_name
IHhp1: interp_list hp1 ** interp_list hp2 ==> interp_list (hp1 ++ hp2)interp_list (a :: hp1) ** interp_list hp2 ==> interp_list ((a :: hp1) ++ hp2)hmap: hprop_map
hp2: list hprop_nameinterp_list hp2 ==> interp_list hp2hmap: hprop_map
a: hprop_name
hp1, hp2: list hprop_name
IHhp1: interp_list hp1 ** interp_list hp2 ==> interp_list (hp1 ++ hp2)interp_list (a :: hp1) ** interp_list hp2 ==> interp_list ((a :: hp1) ++ hp2)hmap: hprop_map
a: hprop_name
hp1, hp2: list hprop_name
IHhp1: interp_list hp1 ** interp_list hp2 ==> interp_list (hp1 ++ hp2)interp_list (a :: hp1) ** interp_list hp2 ==> interp_list ((a :: hp1) ++ hp2)sep. Qed.hmap: hprop_map
a: hprop_name
hp1, hp2: list hprop_name
IHhp1: interp_list hp1 ** interp_list hp2 ==> interp_list (hp1 ++ hp2)(lookup_hprop a hmap ** interp_list hp1) ** interp_list hp2 ==> lookup_hprop a hmap ** interp_list (hp1 ++ hp2)hmap: hprop_map
hp: Hprophinterp hp ==> interp_list (flatten hp)hmap: hprop_map
hp: Hprophinterp hp ==> interp_list (flatten hp)hmap: hprop_mapemp ==> emphmap: hprop_map
h: hprop_namelookup_hprop h hmap ==> lookup_hprop h hmap ** emphmap: hprop_map
hp1, hp2: Hprop
IHhp1: hinterp hp1 ==> interp_list (flatten hp1)
IHhp2: hinterp hp2 ==> interp_list (flatten hp2)hinterp hp1 ** hinterp hp2 ==> interp_list (flatten hp1 ++ flatten hp2)hmap: hprop_map
h: hprop_namelookup_hprop h hmap ==> lookup_hprop h hmap ** emphmap: hprop_map
hp1, hp2: Hprop
IHhp1: hinterp hp1 ==> interp_list (flatten hp1)
IHhp2: hinterp hp2 ==> interp_list (flatten hp2)hinterp hp1 ** hinterp hp2 ==> interp_list (flatten hp1 ++ flatten hp2)hmap: hprop_map
h: hprop_namelookup_hprop h hmap ==> lookup_hprop h hmaphmap: hprop_map
hp1, hp2: Hprop
IHhp1: hinterp hp1 ==> interp_list (flatten hp1)
IHhp2: hinterp hp2 ==> interp_list (flatten hp2)hinterp hp1 ** hinterp hp2 ==> interp_list (flatten hp1 ++ flatten hp2)hmap: hprop_map
hp1, hp2: Hprop
IHhp1: hinterp hp1 ==> interp_list (flatten hp1)
IHhp2: hinterp hp2 ==> interp_list (flatten hp2)hinterp hp1 ** hinterp hp2 ==> interp_list (flatten hp1 ++ flatten hp2)hmap: hprop_map
hp1, hp2: Hprop
IHhp1: hinterp hp1 ==> interp_list (flatten hp1)
IHhp2: hinterp hp2 ==> interp_list (flatten hp2)?Q ==> interp_list (flatten hp1 ++ flatten hp2)hmap: hprop_map
hp1, hp2: Hprop
IHhp1: hinterp hp1 ==> interp_list (flatten hp1)
IHhp2: hinterp hp2 ==> interp_list (flatten hp2)hinterp hp1 ** hinterp hp2 ==> ?Qeapply himp_split ; eauto. Qed.hmap: hprop_map
hp1, hp2: Hprop
IHhp1: hinterp hp1 ==> interp_list (flatten hp1)
IHhp2: hinterp hp2 ==> interp_list (flatten hp2)hinterp hp1 ** hinterp hp2 ==> interp_list (flatten hp1) ** interp_list (flatten hp2)hmap: hprop_map
n: hprop_name
hp: list hprop_namematch remove_one n hp with | Some hp' => exists hp1 hp2 : list hprop_name, hp = hp1 ++ n :: hp2 /\ hp' = hp1 ++ hp2 | None => True endhmap: hprop_map
n: hprop_name
hp: list hprop_namematch remove_one n hp with | Some hp' => exists hp1 hp2 : list hprop_name, hp = hp1 ++ n :: hp2 /\ hp' = hp1 ++ hp2 | None => True endhmap: hprop_map
n, a: hprop_name
hp: list hprop_name
IHhp: match remove_one n hp with | Some hp' => exists hp1 hp2 : list hprop_name, hp = hp1 ++ n :: hp2 /\ hp' = hp1 ++ hp2 | None => True endmatch (if hprop_name_eq n a then Some hp else match remove_one n hp with | Some hps' => Some (a :: hps') | None => None end) with | Some hp' => exists hp1 hp2 : list hprop_name, a :: hp = hp1 ++ n :: hp2 /\ hp' = hp1 ++ hp2 | None => True endhmap: hprop_map
n, a: hprop_name
hp: list hprop_name
IHhp: match remove_one n hp with | Some hp' => exists hp1 hp2 : list hprop_name, hp = hp1 ++ n :: hp2 /\ hp' = hp1 ++ hp2 | None => True end
e: n = aexists hp1 hp2 : list hprop_name, a :: hp = hp1 ++ n :: hp2 /\ hp = hp1 ++ hp2hmap: hprop_map
n, a: hprop_name
hp: list hprop_name
IHhp: match remove_one n hp with | Some hp' => exists hp1 hp2 : list hprop_name, hp = hp1 ++ n :: hp2 /\ hp' = hp1 ++ hp2 | None => True end
n0: n <> amatch match remove_one n hp with | Some hps' => Some (a :: hps') | None => None end with | Some hp' => exists hp1 hp2 : list hprop_name, a :: hp = hp1 ++ n :: hp2 /\ hp' = hp1 ++ hp2 | None => True endhmap: hprop_map
a: hprop_name
hp: list hprop_name
IHhp: match remove_one a hp with | Some hp' => exists hp1 hp2 : list hprop_name, hp = hp1 ++ a :: hp2 /\ hp' = hp1 ++ hp2 | None => True endexists hp1 hp2 : list hprop_name, a :: hp = hp1 ++ a :: hp2 /\ hp = hp1 ++ hp2hmap: hprop_map
n, a: hprop_name
hp: list hprop_name
IHhp: match remove_one n hp with | Some hp' => exists hp1 hp2 : list hprop_name, hp = hp1 ++ n :: hp2 /\ hp' = hp1 ++ hp2 | None => True end
n0: n <> amatch match remove_one n hp with | Some hps' => Some (a :: hps') | None => None end with | Some hp' => exists hp1 hp2 : list hprop_name, a :: hp = hp1 ++ n :: hp2 /\ hp' = hp1 ++ hp2 | None => True endhmap: hprop_map
a: hprop_name
hp: list hprop_name
IHhp: match remove_one a hp with | Some hp' => exists hp1 hp2 : list hprop_name, hp = hp1 ++ a :: hp2 /\ hp' = hp1 ++ hp2 | None => True endexists hp2 : list hprop_name, a :: hp = nil ++ a :: hp2 /\ hp = nil ++ hp2hmap: hprop_map
n, a: hprop_name
hp: list hprop_name
IHhp: match remove_one n hp with | Some hp' => exists hp1 hp2 : list hprop_name, hp = hp1 ++ n :: hp2 /\ hp' = hp1 ++ hp2 | None => True end
n0: n <> amatch match remove_one n hp with | Some hps' => Some (a :: hps') | None => None end with | Some hp' => exists hp1 hp2 : list hprop_name, a :: hp = hp1 ++ n :: hp2 /\ hp' = hp1 ++ hp2 | None => True endhmap: hprop_map
a: hprop_name
hp: list hprop_name
IHhp: match remove_one a hp with | Some hp' => exists hp1 hp2 : list hprop_name, hp = hp1 ++ a :: hp2 /\ hp' = hp1 ++ hp2 | None => True enda :: hp = nil ++ a :: hp /\ hp = nil ++ hphmap: hprop_map
n, a: hprop_name
hp: list hprop_name
IHhp: match remove_one n hp with | Some hp' => exists hp1 hp2 : list hprop_name, hp = hp1 ++ n :: hp2 /\ hp' = hp1 ++ hp2 | None => True end
n0: n <> amatch match remove_one n hp with | Some hps' => Some (a :: hps') | None => None end with | Some hp' => exists hp1 hp2 : list hprop_name, a :: hp = hp1 ++ n :: hp2 /\ hp' = hp1 ++ hp2 | None => True endhmap: hprop_map
n, a: hprop_name
hp: list hprop_name
IHhp: match remove_one n hp with | Some hp' => exists hp1 hp2 : list hprop_name, hp = hp1 ++ n :: hp2 /\ hp' = hp1 ++ hp2 | None => True end
n0: n <> amatch match remove_one n hp with | Some hps' => Some (a :: hps') | None => None end with | Some hp' => exists hp1 hp2 : list hprop_name, a :: hp = hp1 ++ n :: hp2 /\ hp' = hp1 ++ hp2 | None => True endhmap: hprop_map
n, a: hprop_name
hp, l: list hprop_name
IHhp: exists hp1 hp2 : list hprop_name, hp = hp1 ++ n :: hp2 /\ l = hp1 ++ hp2
n0: n <> aexists hp1 hp2 : list hprop_name, a :: hp = hp1 ++ n :: hp2 /\ a :: l = hp1 ++ hp2hmap: hprop_map
n, a: hprop_name
hp: list hprop_name
IHhp: True
n0: n <> aTruehmap: hprop_map
n, a: hprop_name
hp, l, x, x0: list hprop_name
H: hp = x ++ n :: x0
H0: l = x ++ x0
n0: n <> aexists hp1 hp2 : list hprop_name, a :: hp = hp1 ++ n :: hp2 /\ a :: l = hp1 ++ hp2hmap: hprop_map
n, a: hprop_name
hp: list hprop_name
IHhp: True
n0: n <> aTruehmap: hprop_map
n, a: hprop_name
x, x0: list hprop_name
n0: n <> aexists hp1 hp2 : list hprop_name, a :: x ++ n :: x0 = hp1 ++ n :: hp2 /\ a :: x ++ x0 = hp1 ++ hp2hmap: hprop_map
n, a: hprop_name
hp: list hprop_name
IHhp: True
n0: n <> aTruehmap: hprop_map
n, a: hprop_name
x, x0: list hprop_name
n0: n <> aexists hp2 : list hprop_name, a :: x ++ n :: x0 = (a :: x) ++ n :: hp2 /\ a :: x ++ x0 = (a :: x) ++ hp2hmap: hprop_map
n, a: hprop_name
hp: list hprop_name
IHhp: True
n0: n <> aTruehmap: hprop_map
n, a: hprop_name
x, x0: list hprop_name
n0: n <> aa :: x ++ n :: x0 = (a :: x) ++ n :: x0 /\ a :: x ++ x0 = (a :: x) ++ x0hmap: hprop_map
n, a: hprop_name
hp: list hprop_name
IHhp: True
n0: n <> aTruemysimp. Qed.hmap: hprop_map
n, a: hprop_name
hp: list hprop_name
IHhp: True
n0: n <> aTruehmap: hprop_map
x1, x2: list hprop_nameinterp_list (x1 ++ x2) ==> interp_list (x2 ++ x1)hmap: hprop_map
x1, x2: list hprop_nameinterp_list (x1 ++ x2) ==> interp_list (x2 ++ x1)hmap: hprop_map
x2: list hprop_nameinterp_list x2 ==> interp_list (x2 ++ nil)hmap: hprop_map
a: hprop_name
x1, x2: list hprop_name
IHx1: interp_list (x1 ++ x2) ==> interp_list (x2 ++ x1)lookup_hprop a hmap ** interp_list (x1 ++ x2) ==> interp_list (x2 ++ a :: x1)hmap: hprop_map
x2: list hprop_nameinterp_list x2 ==> interp_list (x2 ++ nil)hmap: hprop_map
a: hprop_name
x1, x2: list hprop_name
IHx1: interp_list (x1 ++ x2) ==> interp_list (x2 ++ x1)lookup_hprop a hmap ** interp_list (x1 ++ x2) ==> interp_list (x2 ++ a :: x1)hmap: hprop_map
x2: list hprop_nameinterp_list x2 ==> interp_list x2hmap: hprop_map
a: hprop_name
x1, x2: list hprop_name
IHx1: interp_list (x1 ++ x2) ==> interp_list (x2 ++ x1)lookup_hprop a hmap ** interp_list (x1 ++ x2) ==> interp_list (x2 ++ a :: x1)hmap: hprop_map
a: hprop_name
x1, x2: list hprop_name
IHx1: interp_list (x1 ++ x2) ==> interp_list (x2 ++ x1)lookup_hprop a hmap ** interp_list (x1 ++ x2) ==> interp_list (x2 ++ a :: x1)hmap: hprop_map
a: hprop_name
x1, x2: list hprop_name
IHx1: interp_list (x1 ++ x2) ==> interp_list (x2 ++ x1)lookup_hprop a hmap ** interp_list (x1 ++ x2) ==> interp_list (x2 ++ a :: x1)hmap: hprop_map
a: hprop_name
x1, x2: list hprop_name
IHx1: interp_list (x1 ++ x2) ==> interp_list (x2 ++ x1)?Q ==> interp_list (x2 ++ a :: x1)hmap: hprop_map
a: hprop_name
x1, x2: list hprop_name
IHx1: interp_list (x1 ++ x2) ==> interp_list (x2 ++ x1)lookup_hprop a hmap ** interp_list (x1 ++ x2) ==> ?Qhmap: hprop_map
a: hprop_name
x1, x2: list hprop_name
IHx1: interp_list (x1 ++ x2) ==> interp_list (x2 ++ x1)lookup_hprop a hmap ** interp_list (x1 ++ x2) ==> interp_list x2 ** interp_list (a :: x1)hmap: hprop_map
a: hprop_name
x1, x2: list hprop_name
IHx1: interp_list (x1 ++ x2) ==> interp_list (x2 ++ x1)lookup_hprop a hmap ** interp_list (x1 ++ x2) ==> interp_list (a :: x1) ** interp_list x2hmap: hprop_map
a: hprop_name
x1, x2: list hprop_name
IHx1: interp_list (x1 ++ x2) ==> interp_list (x2 ++ x1)lookup_hprop a hmap ** interp_list (x1 ++ x2) ==> (lookup_hprop a hmap ** interp_list x1) ** interp_list x2apply interp_list_app. Qed.hmap: hprop_map
a: hprop_name
x1, x2: list hprop_name
IHx1: interp_list (x1 ++ x2) ==> interp_list (x2 ++ x1)interp_list (x1 ++ x2) ==> interp_list x1 ** interp_list x2hmap: hprop_map
n: hprop_name
hp1, hp2: list hprop_namematch remove_one n hp1 with | Some hp1' => match remove_one n hp2 with | Some hp2' => interp_list hp1' ==> interp_list hp2' -> interp_list hp1 ==> interp_list hp2 | None => True end | None => True endhmap: hprop_map
n: hprop_name
hp1, hp2: list hprop_namematch remove_one n hp1 with | Some hp1' => match remove_one n hp2 with | Some hp2' => interp_list hp1' ==> interp_list hp2' -> interp_list hp1 ==> interp_list hp2 | None => True end | None => True endhmap: hprop_map
n: hprop_name
hp1, hp2: list hprop_namematch remove_one n hp1 with | Some hp1' => match remove_one n hp2 with | Some hp2' => interp_list hp1' ==> interp_list hp2' -> interp_list hp1 ==> interp_list hp2 | None => True end | None => True endhmap: hprop_map
n: hprop_name
hp1, hp2: list hprop_namematch remove_one n hp1 with | Some hp' => exists hp2 hp3 : list hprop_name, hp1 = hp2 ++ n :: hp3 /\ hp' = hp2 ++ hp3 | None => True end -> match remove_one n hp1 with | Some hp1' => match remove_one n hp2 with | Some hp2' => interp_list hp1' ==> interp_list hp2' -> interp_list hp1 ==> interp_list hp2 | None => True end | None => True endhmap: hprop_map
n: hprop_name
hp1, hp2: list hprop_namematch remove_one n hp2 with | Some hp' => exists hp1 hp3 : list hprop_name, hp2 = hp1 ++ n :: hp3 /\ hp' = hp1 ++ hp3 | None => True end -> match remove_one n hp1 with | Some hp' => exists hp2 hp3 : list hprop_name, hp1 = hp2 ++ n :: hp3 /\ hp' = hp2 ++ hp3 | None => True end -> match remove_one n hp1 with | Some hp1' => match remove_one n hp2 with | Some hp2' => interp_list hp1' ==> interp_list hp2' -> interp_list hp1 ==> interp_list hp2 | None => True end | None => True endhmap: hprop_map
n: hprop_name
hp1, hp2, l: list hprop_namematch remove_one n hp2 with | Some hp' => exists hp1 hp3 : list hprop_name, hp2 = hp1 ++ n :: hp3 /\ hp' = hp1 ++ hp3 | None => True end -> (exists hp2 hp3 : list hprop_name, hp1 = hp2 ++ n :: hp3 /\ l = hp2 ++ hp3) -> match remove_one n hp2 with | Some hp2' => interp_list l ==> interp_list hp2' -> interp_list hp1 ==> interp_list hp2 | None => True endhmap: hprop_map
n: hprop_name
hp1, hp2, l, l0: list hprop_name(exists hp1 hp3 : list hprop_name, hp2 = hp1 ++ n :: hp3 /\ l0 = hp1 ++ hp3) -> (exists hp2 hp3 : list hprop_name, hp1 = hp2 ++ n :: hp3 /\ l = hp2 ++ hp3) -> interp_list l ==> interp_list l0 -> interp_list hp1 ==> interp_list hp2hmap: hprop_map
n: hprop_name
hp1, hp2, l, l0, x1, x2: list hprop_name
H: hp2 = x1 ++ n :: x2
H3: l0 = x1 ++ x2
x, x0: list hprop_name
H0: hp1 = x ++ n :: x0
H2: l = x ++ x0
H1: interp_list l ==> interp_list l0interp_list hp1 ==> interp_list hp2hmap: hprop_map
n: hprop_name
x1, x2, x, x0: list hprop_name
H1: interp_list (x ++ x0) ==> interp_list (x1 ++ x2)interp_list (x ++ n :: x0) ==> interp_list (x1 ++ n :: x2)hmap: hprop_map
n: hprop_name
x1, x2, x, x0: list hprop_name
H1: interp_list (x ++ x0) ==> interp_list (x1 ++ x2)?Q ==> interp_list (x1 ++ n :: x2)hmap: hprop_map
n: hprop_name
x1, x2, x, x0: list hprop_name
H1: interp_list (x ++ x0) ==> interp_list (x1 ++ x2)interp_list (x ++ n :: x0) ==> ?Qhmap: hprop_map
n: hprop_name
x1, x2, x, x0: list hprop_name
H1: interp_list (x ++ x0) ==> interp_list (x1 ++ x2)interp_list (x ++ n :: x0) ==> interp_list ((n :: x2) ++ x1)hmap: hprop_map
n: hprop_name
x1, x2, x, x0: list hprop_name
H1: interp_list (x ++ x0) ==> interp_list (x1 ++ x2)interp_list (x ++ n :: x0) ==> lookup_hprop n hmap ** interp_list (x2 ++ x1)hmap: hprop_map
n: hprop_name
x1, x2, x, x0: list hprop_name
H1: interp_list (x ++ x0) ==> interp_list (x1 ++ x2)interp_list ((n :: x0) ++ x) ==> lookup_hprop n hmap ** interp_list (x2 ++ x1)hmap: hprop_map
n: hprop_name
x1, x2, x, x0: list hprop_name
H1: interp_list (x ++ x0) ==> interp_list (x1 ++ x2)interp_list (x0 ++ x) ==> interp_list (x2 ++ x1)hmap: hprop_map
n: hprop_name
x1, x2, x, x0: list hprop_name
H1: interp_list (x ++ x0) ==> interp_list (x1 ++ x2)?Q ==> interp_list (x2 ++ x1)hmap: hprop_map
n: hprop_name
x1, x2, x, x0: list hprop_name
H1: interp_list (x ++ x0) ==> interp_list (x1 ++ x2)interp_list (x0 ++ x) ==> ?Qhmap: hprop_map
n: hprop_name
x1, x2, x, x0: list hprop_name
H1: interp_list (x ++ x0) ==> interp_list (x1 ++ x2)interp_list (x0 ++ x) ==> interp_list (x1 ++ x2)hmap: hprop_map
n: hprop_name
x1, x2, x, x0: list hprop_name
H1: interp_list (x ++ x0) ==> interp_list (x1 ++ x2)interp_list (x ++ x0) ==> interp_list (x1 ++ x2)hmap: hprop_map
n: hprop_name
x1, x2, x, x0: list hprop_name
H1: interp_list (x ++ x0) ==> interp_list (x1 ++ x2)interp_list (x0 ++ x) ==> interp_list (x ++ x0)apply interp_list_comm. Qed.hmap: hprop_map
n: hprop_name
x1, x2, x, x0: list hprop_name
H1: interp_list (x ++ x0) ==> interp_list (x1 ++ x2)interp_list (x0 ++ x) ==> interp_list (x ++ x0)hmap: hprop_map
n: hprop_name
hp1, hp2: list hprop_namematch remove_one n hp2 with | Some hp2' => interp_list hp1 ==> interp_list hp2' -> interp_list (n :: hp1) ==> interp_list hp2 | None => True endhmap: hprop_map
n: hprop_name
hp1, hp2: list hprop_namematch remove_one n hp2 with | Some hp2' => interp_list hp1 ==> interp_list hp2' -> interp_list (n :: hp1) ==> interp_list hp2 | None => True endhmap: hprop_map
n: hprop_name
hp1, hp2: list hprop_namematch remove_one n hp2 with | Some hp2' => interp_list hp1 ==> interp_list hp2' -> interp_list (n :: hp1) ==> interp_list hp2 | None => True endhmap: hprop_map
n: hprop_name
hp1, hp2: list hprop_namematch remove_one n (n :: hp1) with | Some hp1' => match remove_one n hp2 with | Some hp2' => interp_list hp1' ==> interp_list hp2' -> interp_list (n :: hp1) ==> interp_list hp2 | None => True end | None => True end -> match remove_one n hp2 with | Some hp2' => interp_list hp1 ==> interp_list hp2' -> interp_list (n :: hp1) ==> interp_list hp2 | None => True enddestruct (hprop_name_eq n n) ; mysimp. Qed.hmap: hprop_map
n: hprop_name
hp1, hp2: list hprop_namematch (if hprop_name_eq n n then Some hp1 else match remove_one n hp1 with | Some hps' => Some (n :: hps') | None => None end) with | Some hp1' => match remove_one n hp2 with | Some hp2' => interp_list hp1' ==> interp_list hp2' -> lookup_hprop n hmap ** interp_list hp1 ==> interp_list hp2 | None => True end | None => True end -> match remove_one n hp2 with | Some hp2' => interp_list hp1 ==> interp_list hp2' -> lookup_hprop n hmap ** interp_list hp1 ==> interp_list hp2 | None => True end
This is the key invariant for our simplify
routine.
hmap: hprop_mapforall hp1 hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2hmap: hprop_mapforall hp1 hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2hmap: hprop_map
hp0, hp2: list hprop_name
H: interp_list hp0 ==> interp_list hp2interp_list (hp0 ++ nil) ==> interp_list hp2hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2: list hprop_name
H: interp_list (fst match remove_one a hp2 with | Some hp2' => simplify hp0 hp1 hp2' | None => simplify (hp0 ++ a :: nil) hp1 hp2 end) ==> interp_list (snd match remove_one a hp2 with | Some hp2' => simplify hp0 hp1 hp2' | None => simplify (hp0 ++ a :: nil) hp1 hp2 end)interp_list (hp0 ++ a :: hp1) ==> interp_list hp2hmap: hprop_map
hp0, hp2: list hprop_name
H: interp_list hp0 ==> interp_list hp2interp_list (hp0 ++ nil) ==> interp_list hp2auto.hmap: hprop_map
hp0, hp2: list hprop_name
H: interp_list hp0 ==> interp_list hp2interp_list hp0 ==> interp_list hp2hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2: list hprop_name
H: interp_list (fst match remove_one a hp2 with | Some hp2' => simplify hp0 hp1 hp2' | None => simplify (hp0 ++ a :: nil) hp1 hp2 end) ==> interp_list (snd match remove_one a hp2 with | Some hp2' => simplify hp0 hp1 hp2' | None => simplify (hp0 ++ a :: nil) hp1 hp2 end)interp_list (hp0 ++ a :: hp1) ==> interp_list hp2hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2: list hprop_name
H: interp_list (fst match remove_one a hp2 with | Some hp2' => simplify hp0 hp1 hp2' | None => simplify (hp0 ++ a :: nil) hp1 hp2 end) ==> interp_list (snd match remove_one a hp2 with | Some hp2' => simplify hp0 hp1 hp2' | None => simplify (hp0 ++ a :: nil) hp1 hp2 end)match remove_one a hp2 with | Some hp2' => interp_list (hp0 ++ hp1) ==> interp_list hp2' -> interp_list (a :: hp0 ++ hp1) ==> interp_list hp2 | None => True end -> interp_list (hp0 ++ a :: hp1) ==> interp_list hp2hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2: list hprop_name
H: interp_list (fst match remove_one a hp2 with | Some hp2' => simplify hp0 hp1 hp2' | None => simplify (hp0 ++ a :: nil) hp1 hp2 end) ==> interp_list (snd match remove_one a hp2 with | Some hp2' => simplify hp0 hp1 hp2' | None => simplify (hp0 ++ a :: nil) hp1 hp2 end)
H0: match remove_one a hp2 with | Some hp2' => interp_list (hp0 ++ hp1) ==> interp_list hp2' -> interp_list (a :: hp0 ++ hp1) ==> interp_list hp2 | None => True endinterp_list (hp0 ++ a :: hp1) ==> interp_list hp2hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2, l: list hprop_name
H: interp_list (fst (simplify hp0 hp1 l)) ==> interp_list (snd (simplify hp0 hp1 l))
H0: interp_list (hp0 ++ hp1) ==> interp_list l -> interp_list (a :: hp0 ++ hp1) ==> interp_list hp2interp_list (hp0 ++ a :: hp1) ==> interp_list hp2hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2: list hprop_name
H: interp_list (fst (simplify (hp0 ++ a :: nil) hp1 hp2)) ==> interp_list (snd (simplify (hp0 ++ a :: nil) hp1 hp2))
H0: Trueinterp_list (hp0 ++ a :: hp1) ==> interp_list hp2hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2, l: list hprop_name
H: interp_list (fst (simplify hp0 hp1 l)) ==> interp_list (snd (simplify hp0 hp1 l))
H0: interp_list (hp0 ++ hp1) ==> interp_list l -> interp_list (a :: hp0 ++ hp1) ==> interp_list hp2?Q ==> interp_list hp2hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2, l: list hprop_name
H: interp_list (fst (simplify hp0 hp1 l)) ==> interp_list (snd (simplify hp0 hp1 l))
H0: interp_list (hp0 ++ hp1) ==> interp_list l -> interp_list (a :: hp0 ++ hp1) ==> interp_list hp2interp_list (hp0 ++ a :: hp1) ==> ?Qhmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2: list hprop_name
H: interp_list (fst (simplify (hp0 ++ a :: nil) hp1 hp2)) ==> interp_list (snd (simplify (hp0 ++ a :: nil) hp1 hp2))
H0: Trueinterp_list (hp0 ++ a :: hp1) ==> interp_list hp2hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2, l: list hprop_name
H: interp_list (fst (simplify hp0 hp1 l)) ==> interp_list (snd (simplify hp0 hp1 l))
H0: interp_list (hp0 ++ hp1) ==> interp_list l -> interp_list (a :: hp0 ++ hp1) ==> interp_list hp2interp_list (hp0 ++ hp1) ==> interp_list lhmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2, l: list hprop_name
H: interp_list (fst (simplify hp0 hp1 l)) ==> interp_list (snd (simplify hp0 hp1 l))
H0: interp_list (hp0 ++ hp1) ==> interp_list l -> interp_list (a :: hp0 ++ hp1) ==> interp_list hp2interp_list (hp0 ++ a :: hp1) ==> interp_list (a :: hp0 ++ hp1)hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2: list hprop_name
H: interp_list (fst (simplify (hp0 ++ a :: nil) hp1 hp2)) ==> interp_list (snd (simplify (hp0 ++ a :: nil) hp1 hp2))
H0: Trueinterp_list (hp0 ++ a :: hp1) ==> interp_list hp2hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2, l: list hprop_name
H: interp_list (fst (simplify hp0 hp1 l)) ==> interp_list (snd (simplify hp0 hp1 l))
H0: interp_list (hp0 ++ hp1) ==> interp_list l -> interp_list (a :: hp0 ++ hp1) ==> interp_list hp2interp_list (fst (simplify hp0 hp1 l)) ==> interp_list (snd (simplify hp0 hp1 l))hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2, l: list hprop_name
H: interp_list (fst (simplify hp0 hp1 l)) ==> interp_list (snd (simplify hp0 hp1 l))
H0: interp_list (hp0 ++ hp1) ==> interp_list l -> interp_list (a :: hp0 ++ hp1) ==> interp_list hp2interp_list (hp0 ++ a :: hp1) ==> interp_list (a :: hp0 ++ hp1)hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2: list hprop_name
H: interp_list (fst (simplify (hp0 ++ a :: nil) hp1 hp2)) ==> interp_list (snd (simplify (hp0 ++ a :: nil) hp1 hp2))
H0: Trueinterp_list (hp0 ++ a :: hp1) ==> interp_list hp2hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2, l: list hprop_name
H: interp_list (fst (simplify hp0 hp1 l)) ==> interp_list (snd (simplify hp0 hp1 l))
H0: interp_list (hp0 ++ hp1) ==> interp_list l -> interp_list (a :: hp0 ++ hp1) ==> interp_list hp2interp_list (hp0 ++ a :: hp1) ==> interp_list (a :: hp0 ++ hp1)hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2: list hprop_name
H: interp_list (fst (simplify (hp0 ++ a :: nil) hp1 hp2)) ==> interp_list (snd (simplify (hp0 ++ a :: nil) hp1 hp2))
H0: Trueinterp_list (hp0 ++ a :: hp1) ==> interp_list hp2hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2, l: list hprop_name
H: interp_list (fst (simplify hp0 hp1 l)) ==> interp_list (snd (simplify hp0 hp1 l))
H0: interp_list (hp0 ++ hp1) ==> interp_list l -> interp_list (a :: hp0 ++ hp1) ==> interp_list hp2interp_list ((a :: hp1) ++ hp0) ==> interp_list (a :: hp0 ++ hp1)hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2, l: list hprop_name
H: interp_list (fst (simplify hp0 hp1 l)) ==> interp_list (snd (simplify hp0 hp1 l))
H0: interp_list (hp0 ++ hp1) ==> interp_list l -> interp_list (a :: hp0 ++ hp1) ==> interp_list hp2interp_list (hp0 ++ a :: hp1) ==> interp_list ((a :: hp1) ++ hp0)hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2: list hprop_name
H: interp_list (fst (simplify (hp0 ++ a :: nil) hp1 hp2)) ==> interp_list (snd (simplify (hp0 ++ a :: nil) hp1 hp2))
H0: Trueinterp_list (hp0 ++ a :: hp1) ==> interp_list hp2hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2, l: list hprop_name
H: interp_list (fst (simplify hp0 hp1 l)) ==> interp_list (snd (simplify hp0 hp1 l))
H0: interp_list (hp0 ++ hp1) ==> interp_list l -> lookup_hprop a hmap ** interp_list (hp0 ++ hp1) ==> interp_list hp2interp_list (hp1 ++ hp0) ==> interp_list (hp0 ++ hp1)hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2, l: list hprop_name
H: interp_list (fst (simplify hp0 hp1 l)) ==> interp_list (snd (simplify hp0 hp1 l))
H0: interp_list (hp0 ++ hp1) ==> interp_list l -> interp_list (a :: hp0 ++ hp1) ==> interp_list hp2interp_list (hp0 ++ a :: hp1) ==> interp_list ((a :: hp1) ++ hp0)hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2: list hprop_name
H: interp_list (fst (simplify (hp0 ++ a :: nil) hp1 hp2)) ==> interp_list (snd (simplify (hp0 ++ a :: nil) hp1 hp2))
H0: Trueinterp_list (hp0 ++ a :: hp1) ==> interp_list hp2hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2, l: list hprop_name
H: interp_list (fst (simplify hp0 hp1 l)) ==> interp_list (snd (simplify hp0 hp1 l))
H0: interp_list (hp0 ++ hp1) ==> interp_list l -> interp_list (a :: hp0 ++ hp1) ==> interp_list hp2interp_list (hp0 ++ a :: hp1) ==> interp_list ((a :: hp1) ++ hp0)hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2: list hprop_name
H: interp_list (fst (simplify (hp0 ++ a :: nil) hp1 hp2)) ==> interp_list (snd (simplify (hp0 ++ a :: nil) hp1 hp2))
H0: Trueinterp_list (hp0 ++ a :: hp1) ==> interp_list hp2generalize (IHhp1 (hp0 ++ (a::nil)) hp2); rewrite <- app_assoc; simpl; auto. Qed.hmap: hprop_map
a: hprop_name
hp1: list hprop_name
IHhp1: forall hp0 hp2 : list hprop_name, interp_list (fst (simplify hp0 hp1 hp2)) ==> interp_list (snd (simplify hp0 hp1 hp2)) -> interp_list (hp0 ++ hp1) ==> interp_list hp2
hp0, hp2: list hprop_name
H: interp_list (fst (simplify (hp0 ++ a :: nil) hp1 hp2)) ==> interp_list (snd (simplify (hp0 ++ a :: nil) hp1 hp2))
H0: Trueinterp_list (hp0 ++ a :: hp1) ==> interp_list hp2hmap: hprop_map
hp: Hpropinterp_list (flatten hp) ==> hinterp hphmap: hprop_map
hp: Hpropinterp_list (flatten hp) ==> hinterp hphmap: hprop_mapemp ==> emphmap: hprop_map
h: hprop_namelookup_hprop h hmap ** emp ==> lookup_hprop h hmaphmap: hprop_map
hp1, hp2: Hprop
IHhp1: interp_list (flatten hp1) ==> hinterp hp1
IHhp2: interp_list (flatten hp2) ==> hinterp hp2interp_list (flatten hp1 ++ flatten hp2) ==> hinterp hp1 ** hinterp hp2hmap: hprop_map
h: hprop_namelookup_hprop h hmap ** emp ==> lookup_hprop h hmaphmap: hprop_map
hp1, hp2: Hprop
IHhp1: interp_list (flatten hp1) ==> hinterp hp1
IHhp2: interp_list (flatten hp2) ==> hinterp hp2interp_list (flatten hp1 ++ flatten hp2) ==> hinterp hp1 ** hinterp hp2hmap: hprop_map
h: hprop_nameemp ** lookup_hprop h hmap ==> lookup_hprop h hmaphmap: hprop_map
hp1, hp2: Hprop
IHhp1: interp_list (flatten hp1) ==> hinterp hp1
IHhp2: interp_list (flatten hp2) ==> hinterp hp2interp_list (flatten hp1 ++ flatten hp2) ==> hinterp hp1 ** hinterp hp2hmap: hprop_map
h: hprop_namelookup_hprop h hmap ==> lookup_hprop h hmaphmap: hprop_map
hp1, hp2: Hprop
IHhp1: interp_list (flatten hp1) ==> hinterp hp1
IHhp2: interp_list (flatten hp2) ==> hinterp hp2interp_list (flatten hp1 ++ flatten hp2) ==> hinterp hp1 ** hinterp hp2hmap: hprop_map
hp1, hp2: Hprop
IHhp1: interp_list (flatten hp1) ==> hinterp hp1
IHhp2: interp_list (flatten hp2) ==> hinterp hp2interp_list (flatten hp1 ++ flatten hp2) ==> hinterp hp1 ** hinterp hp2hmap: hprop_map
hp1, hp2: Hprop
IHhp1: interp_list (flatten hp1) ==> hinterp hp1
IHhp2: interp_list (flatten hp2) ==> hinterp hp2interp_list (flatten hp1) ** interp_list (flatten hp2) ==> hinterp hp1 ** hinterp hp2hmap: hprop_map
hp1, hp2: Hprop
IHhp1: interp_list (flatten hp1) ==> hinterp hp1
IHhp2: interp_list (flatten hp2) ==> hinterp hp2interp_list (flatten hp1 ++ flatten hp2) ==> interp_list (flatten hp1) ** interp_list (flatten hp2)apply interp_list_app. Qed.hmap: hprop_map
hp1, hp2: Hprop
IHhp1: interp_list (flatten hp1) ==> hinterp hp1
IHhp2: interp_list (flatten hp2) ==> hinterp hp2interp_list (flatten hp1 ++ flatten hp2) ==> interp_list (flatten hp1) ** interp_list (flatten hp2)
This is the proof that the cross_off
algorithm is correct. That is, if we can prove that the resulting implication holds after crossing off, then the original implication holds.
hmap: hprop_mapforall hp1 hp2 : Hprop, let (hp1', hp2') := cross_off hp1 hp2 in hinterp hp1' ==> hinterp hp2' -> hinterp hp1 ==> hinterp hp2hmap: hprop_mapforall hp1 hp2 : Hprop, let (hp1', hp2') := cross_off hp1 hp2 in hinterp hp1' ==> hinterp hp2' -> hinterp hp1 ==> hinterp hp2hmap: hprop_map
hp1, hp2: Hproplet (hp1', hp2') := cross_off hp1 hp2 in hinterp hp1' ==> hinterp hp2' -> hinterp hp1 ==> hinterp hp2hmap: hprop_map
hp1, hp2: Hprop
e: (Hprop * Hprop)%type
Heqe: e = cross_off hp1 hp2let (hp1', hp2') := e in hinterp hp1' ==> hinterp hp2' -> hinterp hp1 ==> hinterp hp2hmap: hprop_map
hp1, hp2: Hprop
e: (Hprop * Hprop)%type
Heqe: e = (let (hp1', hp2') := simplify nil (flatten hp1) (flatten hp2) in (collapse hp1', collapse hp2'))let (hp1', hp2') := e in hinterp hp1' ==> hinterp hp2' -> hinterp hp1 ==> hinterp hp2hmap: hprop_map
hp1, hp2, h, h0: Hprop
Heqe: (h, h0) = (let (hp1', hp2') := simplify nil (flatten hp1) (flatten hp2) in (collapse hp1', collapse hp2'))hinterp h ==> hinterp h0 -> hinterp hp1 ==> hinterp hp2hmap: hprop_map
hp1, hp2, h, h0: Hprop
Heqe: (h, h0) = (let (hp1', hp2') := simplify nil (flatten hp1) (flatten hp2) in (collapse hp1', collapse hp2'))
H: hinterp h ==> hinterp h0hinterp hp1 ==> hinterp hp2hmap: hprop_map
hp1, hp2, h, h0: Hprop
Heqe: (h, h0) = (let (hp1', hp2') := simplify nil (flatten hp1) (flatten hp2) in (collapse hp1', collapse hp2'))
H: hinterp h ==> hinterp h0interp_list (flatten hp1) ==> hinterp hp2hmap: hprop_map
hp1, hp2, h, h0: Hprop
Heqe: (h, h0) = (let (hp1', hp2') := simplify nil (flatten hp1) (flatten hp2) in (collapse hp1', collapse hp2'))
H: hinterp h ==> hinterp h0(interp_list (fst (simplify nil (flatten hp1) (flatten hp2))) ==> interp_list (snd (simplify nil (flatten hp1) (flatten hp2))) -> interp_list (nil ++ flatten hp1) ==> interp_list (flatten hp2)) -> interp_list (flatten hp1) ==> hinterp hp2hmap: hprop_map
hp1, hp2, h, h0: Hprop
Heqe: (h, h0) = (let (hp1', hp2') := simplify nil (flatten hp1) (flatten hp2) in (collapse hp1', collapse hp2'))
H: hinterp h ==> hinterp h0
H0: interp_list (fst (simplify nil (flatten hp1) (flatten hp2))) ==> interp_list (snd (simplify nil (flatten hp1) (flatten hp2))) -> interp_list (nil ++ flatten hp1) ==> interp_list (flatten hp2)interp_list (flatten hp1) ==> hinterp hp2hmap: hprop_map
hp1, hp2, h, h0: Hprop
Heqe: (h, h0) = (let (hp1', hp2') := simplify nil (flatten hp1) (flatten hp2) in (collapse hp1', collapse hp2'))
H: hinterp h ==> hinterp h0
H0: interp_list (fst (simplify nil (flatten hp1) (flatten hp2))) ==> interp_list (snd (simplify nil (flatten hp1) (flatten hp2))) -> interp_list (nil ++ flatten hp1) ==> interp_list (flatten hp2)?Q ==> hinterp hp2hmap: hprop_map
hp1, hp2, h, h0: Hprop
Heqe: (h, h0) = (let (hp1', hp2') := simplify nil (flatten hp1) (flatten hp2) in (collapse hp1', collapse hp2'))
H: hinterp h ==> hinterp h0
H0: interp_list (fst (simplify nil (flatten hp1) (flatten hp2))) ==> interp_list (snd (simplify nil (flatten hp1) (flatten hp2))) -> interp_list (nil ++ flatten hp1) ==> interp_list (flatten hp2)interp_list (flatten hp1) ==> ?Qhmap: hprop_map
hp1, hp2, h, h0: Hprop
Heqe: (h, h0) = (let (hp1', hp2') := simplify nil (flatten hp1) (flatten hp2) in (collapse hp1', collapse hp2'))
H: hinterp h ==> hinterp h0
H0: interp_list (fst (simplify nil (flatten hp1) (flatten hp2))) ==> interp_list (snd (simplify nil (flatten hp1) (flatten hp2))) -> interp_list (nil ++ flatten hp1) ==> interp_list (flatten hp2)interp_list (flatten hp1) ==> interp_list (flatten hp2)hmap: hprop_map
hp1, hp2, h, h0: Hprop
Heqe: (h, h0) = (let (hp1', hp2') := simplify nil (flatten hp1) (flatten hp2) in (collapse hp1', collapse hp2'))
H: hinterp h ==> hinterp h0
H0: interp_list (fst (simplify nil (flatten hp1) (flatten hp2))) ==> interp_list (snd (simplify nil (flatten hp1) (flatten hp2))) -> interp_list (flatten hp1) ==> interp_list (flatten hp2)interp_list (flatten hp1) ==> interp_list (flatten hp2)hmap: hprop_map
hp1, hp2, h, h0: Hprop
Heqe: (h, h0) = (let (hp1', hp2') := simplify nil (flatten hp1) (flatten hp2) in (collapse hp1', collapse hp2'))
H: hinterp h ==> hinterp h0
H0: interp_list (fst (simplify nil (flatten hp1) (flatten hp2))) ==> interp_list (snd (simplify nil (flatten hp1) (flatten hp2))) -> interp_list (flatten hp1) ==> interp_list (flatten hp2)interp_list (fst (simplify nil (flatten hp1) (flatten hp2))) ==> interp_list (snd (simplify nil (flatten hp1) (flatten hp2)))hmap: hprop_map
hp1, hp2, h, h0: Hprop
l, l0: list hprop_name
Heqe: (h, h0) = (collapse l, collapse l0)
H: hinterp h ==> hinterp h0
H0: interp_list (fst (l, l0)) ==> interp_list (snd (l, l0)) -> interp_list (flatten hp1) ==> interp_list (flatten hp2)interp_list (fst (l, l0)) ==> interp_list (snd (l, l0))hmap: hprop_map
hp1, hp2, h, h0: Hprop
l, l0: list hprop_name
Heqe: (h, h0) = (collapse l, collapse l0)
H: hinterp h ==> hinterp h0
H0: interp_list l ==> interp_list l0 -> interp_list (flatten hp1) ==> interp_list (flatten hp2)interp_list l ==> interp_list l0hmap: hprop_map
hp1, hp2, h, h0: Hprop
l, l0: list hprop_name
Heqe: (h, h0) = (collapse l, collapse l0)
H: hinterp h ==> hinterp h0
H0: interp_list l ==> interp_list l0 -> interp_list (flatten hp1) ==> interp_list (flatten hp2)hinterp (collapse l) ==> hinterp (collapse l0)auto. Qed. End HINTERP.hmap: hprop_map
hp1, hp2: Hprop
l, l0: list hprop_name
Heqe: (collapse l, collapse l0) = (collapse l, collapse l0)
H: hinterp (collapse l) ==> hinterp (collapse l0)
H0: interp_list l ==> interp_list l0 -> interp_list (flatten hp1) ==> interp_list (flatten hp2)hinterp (collapse l) ==> hinterp (collapse l0)
Here is an example using the reflection. We have a large implication to prove that demands many re-associations, commutations, etc. So we first build an hmap
from numbers to hprop
, then we build syntax that parallels the predicates, being sure to use the right atoms according to the hmap
, and finally we invoke the cross_off_corr
lemma to get a much simplified routine. In this case, the simplification gets everything down to emp ==> emp
.
P, Q, R, S, T: hpropQ ** (P ** R) ** S ** T ==> T ** (S ** P) ** Q ** RP, Q, R, S, T: hpropQ ** (P ** R) ** S ** T ==> T ** (S ** P) ** Q ** RP, Q, R, S, T: hprop(let (hp1', hp2') := cross_off (Atom 1 # (Atom 0 # Atom 2) # Atom 3 # Atom 4) (Atom 4 # (Atom 3 # Atom 0) # Atom 1 # Atom 2) in hinterp ((0, P) :: (1, Q) :: (2, R) :: (3, S) :: (4, T) :: nil) hp1' ==> hinterp ((0, P) :: (1, Q) :: (2, R) :: (3, S) :: (4, T) :: nil) hp2' -> hinterp ((0, P) :: (1, Q) :: (2, R) :: (3, S) :: (4, T) :: nil) (Atom 1 # (Atom 0 # Atom 2) # Atom 3 # Atom 4) ==> hinterp ((0, P) :: (1, Q) :: (2, R) :: (3, S) :: (4, T) :: nil) (Atom 4 # (Atom 3 # Atom 0) # Atom 1 # Atom 2)) -> Q ** (P ** R) ** S ** T ==> T ** (S ** P) ** Q ** RP, Q, R, S, T: hprop(emp ==> emp -> Q ** (P ** R) ** S ** T ==> T ** (S ** P) ** Q ** R) -> Q ** (P ** R) ** S ** T ==> T ** (S ** P) ** Q ** RP, Q, R, S, T: hprop
H: emp ==> emp -> Q ** (P ** R) ** S ** T ==> T ** (S ** P) ** Q ** RQ ** (P ** R) ** S ** T ==> T ** (S ** P) ** Q ** Rsep. Qed. Ltac lookup_name term map := match map with | (?n,?P)::?rest => match term with | P => constr:(Some (Atom n)) | _ => lookup_name term rest end | _ => constr:(@None Hprop) end. Ltac reflect term map := match term with | star ?P ?Q => match reflect P map with | (?t1,?map1) => match reflect Q map1 with | (?t2, ?map2) => constr:((Star t1 t2, map2)) end end | emp => constr:((Emp, map)) | _ => match lookup_name term map with | Some ?t => constr:((t, map)) | None => let n := constr:(S (List.length map)) in constr:((Atom n, ((n,term)::map))) end end. Ltac cross := match goal with | [ |- ?A ==> ?B ] => let map := constr:(@nil (hprop_name * hprop)) in match reflect A map with | (?t1, ?map1) => match reflect B map1 with | (?t2, ?map2) => apply (cross_off_corr map2 t1 t2) ; simpl end end end.P, Q, R, S, T: hprop
H: emp ==> emp -> Q ** (P ** R) ** S ** T ==> T ** (S ** P) ** Q ** Remp ==> empP, Q, R, S, T: hpropQ ** (P ** R) ** S ** T ==> T ** (S ** P) ** Q ** RP, Q, R, S, T: hpropQ ** (P ** R) ** S ** T ==> T ** (S ** P) ** Q ** Rsep. Defined.P, Q, R, S, T: hpropemp ==> empp, q, r, s: tptr Nat_t(p --> 0 ** q --> 1) ** (r --> 2 ** emp) ** s --> 3 ==> s --> 3 ** q --> 1 ** p --> 0 ** r --> 2p, q, r, s: tptr Nat_t(p --> 0 ** q --> 1) ** (r --> 2 ** emp) ** s --> 3 ==> s --> 3 ** q --> 1 ** p --> 0 ** r --> 2p, q, r, s: tptr Nat_t(p --> 0 ** q --> 1) ** (r --> 2 ** emp) ** s --> 3 ==> s --> 3 ** q --> 1 ** p --> 0 ** r --> 2sep. Defined. End FunctionalSepIMP.p, q, r, s: tptr Nat_temp ==> emp