mirror of
https://github.com/achlipala/frap.git
synced 2024-12-01 00:26:18 +00:00
Connecting: reading heads
This commit is contained in:
parent
d537e28266
commit
daebf21dc0
1 changed files with 139 additions and 99 deletions
238
Connecting.v
238
Connecting.v
|
@ -1060,7 +1060,7 @@ Module DeeplyEmbedded(Import BW : BIT_WIDTH).
|
||||||
}.
|
}.
|
||||||
End fn.
|
End fn.
|
||||||
|
|
||||||
Definition heap := fmap (wrd) (wrd * wrd).
|
Definition heap := fmap wrd wrd.
|
||||||
Definition valuation := fmap var wrd.
|
Definition valuation := fmap var wrd.
|
||||||
|
|
||||||
Inductive eval : heap -> valuation -> exp -> wrd -> Prop :=
|
Inductive eval : heap -> valuation -> exp -> wrd -> Prop :=
|
||||||
|
@ -1073,13 +1073,13 @@ Module DeeplyEmbedded(Import BW : BIT_WIDTH).
|
||||||
eval H V e1 n1
|
eval H V e1 n1
|
||||||
-> eval H V e2 n2
|
-> eval H V e2 n2
|
||||||
-> eval H V (Add e1 e2) (n1 ^+ n2)
|
-> eval H V (Add e1 e2) (n1 ^+ n2)
|
||||||
| VHead : forall H V e1 p ph pt,
|
| VHead : forall H V e1 p ph,
|
||||||
eval H V e1 p
|
eval H V e1 p
|
||||||
-> H $? p = Some (ph, pt)
|
-> H $? p = Some ph
|
||||||
-> eval H V (Head e1) ph
|
-> eval H V (Head e1) ph
|
||||||
| VTail : forall H V e1 p ph pt,
|
| VTail : forall H V e1 p pt,
|
||||||
eval H V e1 p
|
eval H V e1 p
|
||||||
-> H $? p = Some (ph, pt)
|
-> H $? (p ^+ ^1) = Some pt
|
||||||
-> eval H V (Tail e1) pt
|
-> eval H V (Tail e1) pt
|
||||||
| VNotNull : forall H V e1 p,
|
| VNotNull : forall H V e1 p,
|
||||||
eval H V e1 p
|
eval H V e1 p
|
||||||
|
@ -1089,16 +1089,16 @@ Module DeeplyEmbedded(Import BW : BIT_WIDTH).
|
||||||
| StAssign : forall H V x e v,
|
| StAssign : forall H V x e v,
|
||||||
eval H V e v
|
eval H V e v
|
||||||
-> step (H, V, Assign x e) (H, V $+ (x, v), Skip)
|
-> step (H, V, Assign x e) (H, V $+ (x, v), Skip)
|
||||||
| StWriteHead : forall H V x e a v hv tv,
|
| StWriteHead : forall H V x e a v hv,
|
||||||
V $? x = Some a
|
V $? x = Some a
|
||||||
-> eval H V e v
|
-> eval H V e v
|
||||||
-> H $? a = Some (hv, tv)
|
-> H $? a = Some hv
|
||||||
-> step (H, V, WriteHead x e) (H $+ (a, (v, tv)), V, Skip)
|
-> step (H, V, WriteHead x e) (H $+ (a, v), V, Skip)
|
||||||
| StWriteTail : forall H V x e a v hv tv,
|
| StWriteTail : forall H V x e a v tv,
|
||||||
V $? x = Some a
|
V $? x = Some a
|
||||||
-> eval H V e v
|
-> eval H V e v
|
||||||
-> H $? a = Some (hv, tv)
|
-> H $? a = Some tv
|
||||||
-> step (H, V, WriteTail x e) (H $+ (a, (hv, v)), V, Skip)
|
-> step (H, V, WriteTail x e) (H $+ (a, v), V, Skip)
|
||||||
| StSeq1 : forall H V s2,
|
| StSeq1 : forall H V s2,
|
||||||
step (H, V, Seq Skip s2) (H, V, s2)
|
step (H, V, Seq Skip s2) (H, V, s2)
|
||||||
| StSeq2 : forall H V s1 s2 H' V' s1',
|
| StSeq2 : forall H V s1 s2 H' V' s1',
|
||||||
|
@ -1167,17 +1167,18 @@ Module MixedToDeep(Import BW : BIT_WIDTH).
|
||||||
| TrDone : forall V (v : RT) s,
|
| TrDone : forall V (v : RT) s,
|
||||||
translate_return V v s
|
translate_return V v s
|
||||||
-> translate translate_return V (Return v) s
|
-> translate translate_return V (Return v) s
|
||||||
| TrAssign : forall V (B : Set) (v : wrd) (c : wrd -> cmd B) e x s1,
|
| TrAssign : forall V B (v : wrd) (c : wrd -> cmd B) e x s1,
|
||||||
translate_exp V v e
|
translate_exp V v e
|
||||||
-> (forall w, translate translate_return (V $+ (x, w)) (c w) s1)
|
-> (forall w, translate translate_return (V $+ (x, w)) (c w) s1)
|
||||||
-> translate translate_return V (Bind (Return v) c) (Seq (Assign x e) s1)
|
-> translate translate_return V (Bind (Return v) c) (Seq (Assign x e) s1)
|
||||||
| TrAssigned : forall V (B : Set) (v : wrd) (c : wrd -> cmd B) x s1,
|
| TrAssigned : forall V B (v : wrd) (c : wrd -> cmd B) x s1,
|
||||||
V $? x = Some v
|
V $? x = Some v
|
||||||
-> translate translate_return (V $+ (x, v)) (c v) s1
|
-> translate translate_return V (c v) s1
|
||||||
-> translate translate_return V (Bind (Return v) c) (Seq Skip s1).
|
-> translate translate_return V (Bind (Return v) c) (Seq Skip s1)
|
||||||
|
| TrReadHead : forall V B (a : wrd) (c : wrd -> cmd B) e x s1,
|
||||||
Example adder (a b c : wrd) :=
|
translate_exp V a e
|
||||||
Bind (Return (a ^+ b)) (fun ab => Return (ab ^+ c)).
|
-> (forall w, translate translate_return (V $+ (x, w)) (c w) s1)
|
||||||
|
-> translate translate_return V (Bind (Read a) c) (Seq (Assign x (Head e)) s1).
|
||||||
|
|
||||||
Ltac freshFor vm k :=
|
Ltac freshFor vm k :=
|
||||||
let rec keepTrying x :=
|
let rec keepTrying x :=
|
||||||
|
@ -1199,8 +1200,14 @@ Module MixedToDeep(Import BW : BIT_WIDTH).
|
||||||
| [ |- translate _ ?V (Bind (Return _) _) _ ] =>
|
| [ |- translate _ ?V (Bind (Return _) _) _ ] =>
|
||||||
freshFor V ltac:(fun y =>
|
freshFor V ltac:(fun y =>
|
||||||
eapply TrAssign with (x := y); [ | intro ])
|
eapply TrAssign with (x := y); [ | intro ])
|
||||||
|
| [ |- translate _ ?V (Bind (Read _) _) _ ] =>
|
||||||
|
freshFor V ltac:(fun y =>
|
||||||
|
eapply TrReadHead with (x := y); [ | intro ])
|
||||||
end.
|
end.
|
||||||
|
|
||||||
|
Example adder (a b c : wrd) :=
|
||||||
|
Bind (Return (a ^+ b)) (fun ab => Return (ab ^+ c)).
|
||||||
|
|
||||||
Lemma translate_adder : sig (fun s =>
|
Lemma translate_adder : sig (fun s =>
|
||||||
forall a b c, translate (translate_result "result") ($0 $+ ("a", a) $+ ("b", b) $+ ("c", c)) (adder a b c) s).
|
forall a b c, translate (translate_result "result") ($0 $+ ("a", a) $+ ("b", b) $+ ("c", c)) (adder a b c) s).
|
||||||
Proof.
|
Proof.
|
||||||
|
@ -1211,18 +1218,23 @@ Module MixedToDeep(Import BW : BIT_WIDTH).
|
||||||
|
|
||||||
Definition adder_compiled := Eval simpl in proj1_sig translate_adder.
|
Definition adder_compiled := Eval simpl in proj1_sig translate_adder.
|
||||||
|
|
||||||
Record heaps_compat (H : DE.heap) (h : ME.heap) : Prop := {
|
Example reader (p1 p2 : wrd) :=
|
||||||
DeepToMixedHd : forall a v1 v2, H $? a = Some (v1, v2) -> h $? a = Some v1;
|
Bind (Read p1) (fun v1 => Bind (Read p2) (fun v2 => Return (v1 ^+ v2))).
|
||||||
DeepToMixedTl : forall a v1 v2, H $? a = Some (v1, v2) -> h $? (a ^+ ^1) = Some v2;
|
|
||||||
MixedToDeep : forall a v1 v2, h $? a = Some v1 -> h $? (a ^+ ^1) = Some v2
|
Lemma translate_reader : sig (fun s =>
|
||||||
-> H $? a = Some (v1, v2)
|
forall p1 p2, translate (translate_result "result") ($0 $+ ("p1", p1) $+ ("p2", p2)) (reader p1 p2) s).
|
||||||
}.
|
Proof.
|
||||||
|
eexists; simplify.
|
||||||
|
unfold reader.
|
||||||
|
repeat translate.
|
||||||
|
Defined.
|
||||||
|
|
||||||
|
Definition reader_compiled := Eval simpl in proj1_sig translate_reader.
|
||||||
|
|
||||||
Inductive translated : forall {A}, DE.heap * valuation * stmt -> ME.heap * cmd A -> Prop :=
|
Inductive translated : forall {A}, DE.heap * valuation * stmt -> ME.heap * cmd A -> Prop :=
|
||||||
| Translated : forall A H h V s (c : cmd A),
|
| Translated : forall A H V s (c : cmd A),
|
||||||
translate (translate_result "result") V c s
|
translate (translate_result "result") V c s
|
||||||
-> heaps_compat H h
|
-> translated (H, V, s) (H, c).
|
||||||
-> translated (H, V, s) (h, c).
|
|
||||||
|
|
||||||
Lemma eval_translate : forall H V e v,
|
Lemma eval_translate : forall H V e v,
|
||||||
eval H V e v
|
eval H V e v
|
||||||
|
@ -1253,75 +1265,57 @@ Module MixedToDeep(Import BW : BIT_WIDTH).
|
||||||
eauto.
|
eauto.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Lemma step_translate : forall H V s H' V' s',
|
Lemma step_translate : forall out V (c : cmd wrd) s,
|
||||||
DE.step (H, V, s) (H', V', s')
|
translate (translate_result out) V c s
|
||||||
-> forall h (c : cmd wrd) out,
|
-> forall H H' V' s',
|
||||||
translate (translate_result out) V c s
|
DE.step (H, V, s) (H', V', s')
|
||||||
-> heaps_compat H h
|
-> exists c', ME.step^* (H, c) (H', c')
|
||||||
-> exists c' h', ME.step^* (h, c) (h', c')
|
/\ translate (translate_result out) V' c' s'.
|
||||||
/\ translate (translate_result out) V' c' s'
|
|
||||||
/\ heaps_compat H' h'.
|
|
||||||
Proof.
|
Proof.
|
||||||
induct 1; invert 1; simplify.
|
induct 1.
|
||||||
|
|
||||||
invert H3.
|
invert H; invert 1; simplify.
|
||||||
apply inj_pair2 in H1; subst.
|
|
||||||
eapply eval_translate in H4; eauto; subst.
|
eapply eval_translate in H0; eauto; subst.
|
||||||
do 2 eexists; propositional.
|
do 2 eexists; propositional.
|
||||||
eauto.
|
eauto.
|
||||||
apply TrDone; apply TrReturned; simplify; auto.
|
apply TrDone; apply TrReturned; simplify; auto.
|
||||||
assumption.
|
|
||||||
|
|
||||||
|
invert 1; simplify.
|
||||||
invert H6.
|
invert H6.
|
||||||
|
eapply eval_translate in H5; eauto; subst.
|
||||||
|
do 2 eexists; propositional.
|
||||||
|
eauto.
|
||||||
|
econstructor.
|
||||||
|
instantiate (1 := x); simplify.
|
||||||
|
reflexivity.
|
||||||
|
eauto.
|
||||||
|
|
||||||
invert H6.
|
invert 1; simplify.
|
||||||
|
|
||||||
invert H2.
|
|
||||||
|
|
||||||
apply inj_pair2 in H0; subst.
|
|
||||||
do 2 eexists; propositional.
|
do 2 eexists; propositional.
|
||||||
eapply TrcFront.
|
eapply TrcFront.
|
||||||
eauto.
|
eauto.
|
||||||
eauto.
|
eauto.
|
||||||
replace V' with (V' $+ (x, v)).
|
|
||||||
assumption.
|
|
||||||
maps_equal.
|
|
||||||
symmetry; assumption.
|
|
||||||
assumption.
|
assumption.
|
||||||
|
invert H5.
|
||||||
|
|
||||||
invert H4.
|
invert 1.
|
||||||
|
invert H6.
|
||||||
apply inj_pair2 in H2; subst.
|
invert H5.
|
||||||
invert H0.
|
eapply eval_translate in H4; eauto; subst.
|
||||||
do 2 eexists; propositional.
|
|
||||||
eauto.
|
|
||||||
eapply TrAssigned.
|
|
||||||
instantiate (1 := x).
|
|
||||||
simplify.
|
simplify.
|
||||||
eapply eval_translate in H7; eauto.
|
do 2 eexists; propositional.
|
||||||
subst.
|
eapply TrcFront.
|
||||||
reflexivity.
|
eauto.
|
||||||
match goal with
|
eauto.
|
||||||
| [ |- translate _ ?V' _ _ ] => replace V' with (V $+ (x, v)) by (maps_equal; equality)
|
econstructor.
|
||||||
end.
|
instantiate (1 := x); simplify; reflexivity.
|
||||||
eauto.
|
eauto.
|
||||||
assumption.
|
|
||||||
|
|
||||||
invert H0.
|
|
||||||
|
|
||||||
invert H4.
|
|
||||||
|
|
||||||
invert H3.
|
|
||||||
|
|
||||||
invert H4.
|
|
||||||
|
|
||||||
invert H3.
|
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Theorem translated_simulates : forall H V c h s,
|
Theorem translated_simulates : forall H V c s,
|
||||||
translate (translate_result "result") V c s
|
translate (translate_result "result") V c s
|
||||||
-> heaps_compat H h
|
-> simulates (translated (A := wrd)) (DE.trsys_of H V s) (ME.multistep_trsys_of H c).
|
||||||
-> simulates (translated (A := wrd)) (DE.trsys_of H V s) (ME.multistep_trsys_of h c).
|
|
||||||
Proof.
|
Proof.
|
||||||
constructor; simplify.
|
constructor; simplify.
|
||||||
|
|
||||||
|
@ -1332,16 +1326,15 @@ Module MixedToDeep(Import BW : BIT_WIDTH).
|
||||||
eauto.
|
eauto.
|
||||||
auto.
|
auto.
|
||||||
|
|
||||||
invert H2.
|
invert H1.
|
||||||
apply inj_pair2 in H5; subst.
|
apply inj_pair2 in H4; subst.
|
||||||
cases st1'.
|
cases st1'.
|
||||||
cases p.
|
cases p.
|
||||||
eapply step_translate in H6; eauto.
|
eapply step_translate in H7; eauto.
|
||||||
simp.
|
simp.
|
||||||
eexists; split; [ | eassumption ].
|
eexists; split; [ | eassumption ].
|
||||||
econstructor.
|
econstructor.
|
||||||
assumption.
|
assumption.
|
||||||
eauto.
|
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Hint Constructors eval DE.step.
|
Hint Constructors eval DE.step.
|
||||||
|
@ -1355,10 +1348,11 @@ Module MixedToDeep(Import BW : BIT_WIDTH).
|
||||||
|
|
||||||
Hint Resolve translate_exp_sound.
|
Hint Resolve translate_exp_sound.
|
||||||
|
|
||||||
Lemma not_stuck : forall A h (c : cmd A) h' c',
|
Lemma not_stuck : forall out V (c : cmd wrd) s,
|
||||||
step (h, c) (h', c')
|
translate (translate_result out) V c s
|
||||||
-> forall out V s, translate (translate_result out) V c s
|
-> forall H H' c',
|
||||||
-> forall H, exists p', DE.step (H, V, s) p'.
|
step (H, c) (H', c')
|
||||||
|
-> exists p', DE.step (H, V, s) p'.
|
||||||
Proof.
|
Proof.
|
||||||
induct 1; invert 1; simplify.
|
induct 1; invert 1; simplify.
|
||||||
|
|
||||||
|
@ -1369,20 +1363,27 @@ Module MixedToDeep(Import BW : BIT_WIDTH).
|
||||||
|
|
||||||
eexists.
|
eexists.
|
||||||
econstructor.
|
econstructor.
|
||||||
|
econstructor.
|
||||||
|
eauto.
|
||||||
|
|
||||||
eexists.
|
eexists.
|
||||||
econstructor.
|
econstructor.
|
||||||
econstructor.
|
econstructor.
|
||||||
eauto.
|
eauto.
|
||||||
|
|
||||||
|
apply inj_pair2 in H8; subst.
|
||||||
|
invert H6.
|
||||||
eexists.
|
eexists.
|
||||||
econstructor.
|
econstructor.
|
||||||
|
econstructor.
|
||||||
|
econstructor.
|
||||||
|
eauto.
|
||||||
|
eauto.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Theorem hoare_triple_sound : forall P (c : cmd wrd) Q V s h H,
|
Theorem hoare_triple_sound : forall P (c : cmd wrd) Q V s H,
|
||||||
hoare_triple P c Q
|
hoare_triple P c Q
|
||||||
-> P h
|
-> P H
|
||||||
-> heaps_compat H h
|
|
||||||
-> translate (translate_result "result") V c s
|
-> translate (translate_result "result") V c s
|
||||||
-> V $? "result" = None
|
-> V $? "result" = None
|
||||||
-> invariantFor (DE.trsys_of H V s)
|
-> invariantFor (DE.trsys_of H V s)
|
||||||
|
@ -1394,25 +1395,24 @@ Module MixedToDeep(Import BW : BIT_WIDTH).
|
||||||
eapply invariant_simulates.
|
eapply invariant_simulates.
|
||||||
apply translated_simulates.
|
apply translated_simulates.
|
||||||
eassumption.
|
eassumption.
|
||||||
eassumption.
|
apply invariant_multistepify with (sys := trsys_of H c).
|
||||||
apply invariant_multistepify with (sys := trsys_of h c).
|
|
||||||
eauto using hoare_triple_sound.
|
eauto using hoare_triple_sound.
|
||||||
invert 1; simp.
|
invert 1; simp.
|
||||||
|
|
||||||
cases st2; simplify; subst.
|
cases st2; simplify; subst.
|
||||||
invert H6; simplify.
|
invert H5; simplify.
|
||||||
apply inj_pair2 in H8; subst.
|
apply inj_pair2 in H10; subst.
|
||||||
invert H11.
|
invert H9.
|
||||||
apply inj_pair2 in H6; subst.
|
apply inj_pair2 in H4; subst.
|
||||||
invert H8.
|
invert H6.
|
||||||
right; eexists.
|
right; eexists.
|
||||||
econstructor; eauto.
|
econstructor; eauto.
|
||||||
auto.
|
auto.
|
||||||
|
|
||||||
invert H6; simp.
|
invert H5; simp.
|
||||||
apply inj_pair2 in H8; subst; simplify.
|
apply inj_pair2 in H7; subst; simplify.
|
||||||
cases x.
|
cases x.
|
||||||
eapply not_stuck in H7; eauto.
|
eapply not_stuck in H10; eauto.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Theorem adder_ok : forall a b c,
|
Theorem adder_ok : forall a b c,
|
||||||
|
@ -1441,10 +1441,50 @@ Module MixedToDeep(Import BW : BIT_WIDTH).
|
||||||
eapply hoare_triple_sound.
|
eapply hoare_triple_sound.
|
||||||
apply adder_ok.
|
apply adder_ok.
|
||||||
constructor; auto.
|
constructor; auto.
|
||||||
constructor; simplify; equality.
|
|
||||||
apply (proj2_sig translate_adder).
|
apply (proj2_sig translate_adder).
|
||||||
simplify; equality.
|
simplify; equality.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
|
Theorem reader_ok : forall p1 p2 v1 v2,
|
||||||
|
{{p1 |-> v1 * p2 |-> v2}}
|
||||||
|
reader p1 p2
|
||||||
|
{{r ~> [| r = v1 ^+ v2 |] * p1 |-> v1 * p2 |-> v2}}.
|
||||||
|
Proof.
|
||||||
|
unfold reader.
|
||||||
|
simplify.
|
||||||
|
step.
|
||||||
|
step.
|
||||||
|
simp.
|
||||||
|
step.
|
||||||
|
step.
|
||||||
|
simp.
|
||||||
|
step.
|
||||||
|
cancel.
|
||||||
|
equality.
|
||||||
|
Qed.
|
||||||
|
|
||||||
|
Theorem reader_compiled_ok : forall p1 p2 v1 v2,
|
||||||
|
p1 <> p2
|
||||||
|
-> invariantFor (DE.trsys_of ($0 $+ (p1, v1) $+ (p2, v2)) ($0 $+ ("p1", p1) $+ ("p2", p2)) reader_compiled)
|
||||||
|
(fun p => snd p = Skip
|
||||||
|
\/ exists p', DE.step p p').
|
||||||
|
Proof.
|
||||||
|
simplify.
|
||||||
|
eapply hoare_triple_sound.
|
||||||
|
apply reader_ok.
|
||||||
|
exists ($0 $+ (p1, v1)), ($0 $+ (p2, v2)); propositional.
|
||||||
|
unfold split.
|
||||||
|
maps_equal.
|
||||||
|
rewrite lookup_join2; simplify; auto; sets.
|
||||||
|
rewrite lookup_join1; simplify; auto; sets.
|
||||||
|
rewrite lookup_join2; simplify; auto; sets.
|
||||||
|
unfold disjoint; simplify.
|
||||||
|
cases (weq a p1); simplify; propositional.
|
||||||
|
constructor.
|
||||||
|
constructor.
|
||||||
|
apply (proj2_sig translate_reader).
|
||||||
|
simplify; equality.
|
||||||
|
Qed.
|
||||||
End MixedToDeep.
|
End MixedToDeep.
|
||||||
|
|
||||||
|
|
||||||
|
|
Loading…
Reference in a new issue