mirror of
https://github.com/achlipala/frap.git
synced 2024-11-12 17:17:50 +00:00
Fix a performance bug in model_check
This commit is contained in:
parent
1c88de0f7d
commit
cd2a474d5d
3 changed files with 17 additions and 15 deletions
14
Sets.v
14
Sets.v
|
@ -565,14 +565,16 @@ Ltac quote E env k :=
|
|||
quote' E env k
|
||||
end.
|
||||
|
||||
Ltac sets_cbv := cbv beta iota zeta delta [interp_normal_form normalize_setexpr nth_default
|
||||
setmerge Elements Other nth_error map dedup member beq_nat orb
|
||||
andb included].
|
||||
|
||||
Ltac normalize_set :=
|
||||
match goal with
|
||||
| [ |- context[@union ?A ?X ?Y] ] =>
|
||||
quote (@union A X Y) (@nil A) ltac:(fun e env =>
|
||||
change (@union A X Y) with (interp_setexpr env e);
|
||||
rewrite <- normalize_setexpr_ok;
|
||||
cbv beta iota zeta delta [interp_normal_form normalize_setexpr nth_default
|
||||
setmerge Elements Other nth_error map dedup member beq_nat orb])
|
||||
change (@union A X Y) with (interp_setexpr env e));
|
||||
rewrite <- normalize_setexpr_ok; sets_cbv
|
||||
end.
|
||||
|
||||
Ltac compare_sets :=
|
||||
|
@ -582,7 +584,7 @@ Ltac compare_sets :=
|
|||
| ?A -> _ =>
|
||||
quote X (@nil A) ltac:(fun x env =>
|
||||
quote Y env ltac:(fun y env' =>
|
||||
change (interp_setexpr env' x = interp_setexpr env' y);
|
||||
apply compare_sets; reflexivity))
|
||||
change (interp_setexpr env' x = interp_setexpr env' y)));
|
||||
apply compare_sets; sets_cbv; reflexivity
|
||||
end
|
||||
end.
|
||||
|
|
|
@ -77,7 +77,7 @@ Inductive step : heap * locks * cmd -> heap * locks * cmd -> Prop :=
|
|||
(* To take a lock, it must not be held; and vice versa for releasing a lock. *)
|
||||
| StepLock : forall h l a,
|
||||
~a \in l
|
||||
-> step (h, l, Lock a) (h, l \cup {a}, Return 0)
|
||||
-> step (h, l, Lock a) (h, {a} \cup l, Return 0)
|
||||
| StepUnlock : forall h l a,
|
||||
a \in l
|
||||
-> step (h, l, Unlock a) (h, l \setminus {a}, Return 0).
|
||||
|
@ -603,13 +603,13 @@ Proof.
|
|||
do 2 eexists; propositional.
|
||||
constructor.
|
||||
sets.
|
||||
replace ((l \cup {a}) \cup {a0}) with ((l \cup {a0}) \cup {a}) by sets.
|
||||
replace ({a0, a} \cup l) with ({a} \cup ({a0} \cup l)) by sets.
|
||||
constructor.
|
||||
sets.
|
||||
do 2 eexists; propositional.
|
||||
constructor.
|
||||
sets; propositional.
|
||||
replace (l \cup {a} \setminus {a0}) with ((l \setminus {a0}) \cup {a}) by sets.
|
||||
replace ({a} \cup l \setminus {a0}) with ({a} \cup (l \setminus {a0})) by sets.
|
||||
constructor.
|
||||
sets.
|
||||
|
||||
|
@ -617,7 +617,7 @@ Proof.
|
|||
do 2 eexists; propositional.
|
||||
constructor.
|
||||
sets.
|
||||
replace ((l \setminus {a}) \cup {a0}) with ((l \cup {a0}) \setminus {a}) by sets.
|
||||
replace ({a0} \cup (l \setminus {a})) with (({a0} \cup l) \setminus {a}) by sets.
|
||||
constructor.
|
||||
sets.
|
||||
do 2 eexists; propositional.
|
||||
|
@ -1360,7 +1360,7 @@ Lemma commute_locks : forall c1 c a s h l1' h' l',
|
|||
-> a \in Locks s
|
||||
-> commutes c s
|
||||
-> forall c1', step (h, l1', c1) (h', l', c1')
|
||||
-> step (h, l1' \cup {a}, c1) (h', l' \cup {a}, c1').
|
||||
-> step (h, {a} \cup l1', c1) (h', {a} \cup l', c1').
|
||||
Proof.
|
||||
induct 1; simplify.
|
||||
|
||||
|
@ -1373,12 +1373,12 @@ Proof.
|
|||
invert H1; eauto.
|
||||
|
||||
invert H1.
|
||||
replace ((l1' \cup {l}) \cup {a}) with ((l1' \cup {a}) \cup {l}) by sets.
|
||||
replace ({a} \cup ({l} \cup l1')) with ({l} \cup ({a} \cup l1')) by sets.
|
||||
constructor.
|
||||
sets.
|
||||
|
||||
invert H1.
|
||||
replace ((l1' \setminus {l}) \cup {a}) with ((l1' \cup {a}) \setminus {l}) by sets.
|
||||
replace ({a} \cup (l1' \setminus {l})) with (({a} \cup l1') \setminus {l}) by sets.
|
||||
constructor.
|
||||
sets.
|
||||
|
||||
|
@ -1403,7 +1403,7 @@ Proof.
|
|||
invert H1; eauto.
|
||||
|
||||
invert H1.
|
||||
replace (l1' \cup {l} \setminus {a}) with ((l1' \setminus {a}) \cup {l}) by sets.
|
||||
replace (({l} \cup l1') \setminus {a}) with ({l} \cup (l1' \setminus {a})) by sets.
|
||||
constructor.
|
||||
sets.
|
||||
|
||||
|
|
|
@ -33,7 +33,7 @@ DeepAndShallowEmbeddings.v
|
|||
SepCancel.v
|
||||
SeparationLogic_template.v
|
||||
SeparationLogic.v
|
||||
#SharedMemory.v
|
||||
SharedMemory.v
|
||||
ConcurrentSeparationLogic_template.v
|
||||
ConcurrentSeparationLogic.v
|
||||
MessagesAndRefinement.v
|
||||
|
|
Loading…
Reference in a new issue