mirror of
https://github.com/achlipala/frap.git
synced 2024-11-10 00:07:51 +00:00
Connecting: added a heap relation, but at the moment it could be anything, because no heap-accessing commands are supported
This commit is contained in:
parent
6b3a93a8b2
commit
ca6d577f84
1 changed files with 44 additions and 25 deletions
69
Connecting.v
69
Connecting.v
|
@ -738,8 +738,8 @@ Module MixedEmbedded(Import BW : BIT_WIDTH).
|
||||||
|
|
||||||
Lemma hoare_triple_sound' : forall P {result} (c : cmd result) Q,
|
Lemma hoare_triple_sound' : forall P {result} (c : cmd result) Q,
|
||||||
hoare_triple P c Q
|
hoare_triple P c Q
|
||||||
-> P $0
|
-> forall h, P h
|
||||||
-> invariantFor (trsys_of $0 c)
|
-> invariantFor (trsys_of h c)
|
||||||
(fun p =>
|
(fun p =>
|
||||||
hoare_triple (fun h' => h' = fst p)
|
hoare_triple (fun h' => h' = fst p)
|
||||||
(snd p)
|
(snd p)
|
||||||
|
@ -761,8 +761,8 @@ Module MixedEmbedded(Import BW : BIT_WIDTH).
|
||||||
|
|
||||||
Theorem hoare_triple_sound : forall P {result} (c : cmd result) Q,
|
Theorem hoare_triple_sound : forall P {result} (c : cmd result) Q,
|
||||||
hoare_triple P c Q
|
hoare_triple P c Q
|
||||||
-> P $0
|
-> forall h, P h
|
||||||
-> invariantFor (trsys_of $0 c)
|
-> invariantFor (trsys_of h c)
|
||||||
(fun p => (exists r, snd p = Return r)
|
(fun p => (exists r, snd p = Return r)
|
||||||
\/ (exists p', step p p')).
|
\/ (exists p', step p p')).
|
||||||
Proof.
|
Proof.
|
||||||
|
@ -1205,9 +1205,17 @@ 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 := {
|
||||||
|
DeepToMixedHd : forall a v1 v2, H $? a = Some (v1, v2) -> h $? a = Some v1;
|
||||||
|
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
|
||||||
|
-> H $? a = Some (v1, v2)
|
||||||
|
}.
|
||||||
|
|
||||||
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 h V s (c : cmd A),
|
||||||
translate "result" V c s
|
translate "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,
|
||||||
|
@ -1243,19 +1251,22 @@ Module MixedToDeep(Import BW : BIT_WIDTH).
|
||||||
DE.step (H, V, s) (H', V', s')
|
DE.step (H, V, s) (H', V', s')
|
||||||
-> forall h (c : cmd wrd) out,
|
-> forall h (c : cmd wrd) out,
|
||||||
translate out V c s
|
translate out V c s
|
||||||
-> exists c', ME.step^* (h, c) (h, c')
|
-> heaps_compat H h
|
||||||
/\ translate out V' c' s'.
|
-> exists c' h', ME.step^* (h, c) (h', c')
|
||||||
|
/\ translate out V' c' s'
|
||||||
|
/\ heaps_compat H' h'.
|
||||||
Proof.
|
Proof.
|
||||||
induct 1; invert 1; simplify.
|
induct 1; invert 1; simplify.
|
||||||
|
|
||||||
apply inj_pair2 in H1; subst.
|
apply inj_pair2 in H1; subst.
|
||||||
eapply eval_translate in H5; eauto; subst.
|
eapply eval_translate in H5; eauto; subst.
|
||||||
eexists; propositional.
|
do 2 eexists; propositional.
|
||||||
eauto.
|
eauto.
|
||||||
apply TrReturned; simplify; auto.
|
apply TrReturned; simplify; auto.
|
||||||
|
assumption.
|
||||||
|
|
||||||
apply inj_pair2 in H0; subst.
|
apply inj_pair2 in H0; subst.
|
||||||
eexists; propositional.
|
do 2 eexists; propositional.
|
||||||
eapply TrcFront.
|
eapply TrcFront.
|
||||||
eauto.
|
eauto.
|
||||||
eauto.
|
eauto.
|
||||||
|
@ -1263,10 +1274,11 @@ Module MixedToDeep(Import BW : BIT_WIDTH).
|
||||||
assumption.
|
assumption.
|
||||||
maps_equal.
|
maps_equal.
|
||||||
symmetry; assumption.
|
symmetry; assumption.
|
||||||
|
assumption.
|
||||||
|
|
||||||
apply inj_pair2 in H2; subst.
|
apply inj_pair2 in H2; subst.
|
||||||
invert H0.
|
invert H0.
|
||||||
eexists; propositional.
|
do 2 eexists; propositional.
|
||||||
eauto.
|
eauto.
|
||||||
eapply TrAssigned with (x := x).
|
eapply TrAssigned with (x := x).
|
||||||
eapply eval_translate in H7; eauto.
|
eapply eval_translate in H7; eauto.
|
||||||
|
@ -1277,12 +1289,14 @@ Module MixedToDeep(Import BW : BIT_WIDTH).
|
||||||
| [ |- translate _ ?V' _ _ ] => replace V' with (V $+ (x, v)) by (maps_equal; equality)
|
| [ |- translate _ ?V' _ _ ] => replace V' with (V $+ (x, v)) by (maps_equal; equality)
|
||||||
end.
|
end.
|
||||||
eauto.
|
eauto.
|
||||||
|
assumption.
|
||||||
|
|
||||||
invert H0.
|
invert H0.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Theorem translated_simulates : forall H V c h s,
|
Theorem translated_simulates : forall H V c h s,
|
||||||
translate "result" V c s
|
translate "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.
|
||||||
|
@ -1292,16 +1306,18 @@ Module MixedToDeep(Import BW : BIT_WIDTH).
|
||||||
econstructor.
|
econstructor.
|
||||||
eassumption.
|
eassumption.
|
||||||
eauto.
|
eauto.
|
||||||
|
auto.
|
||||||
|
|
||||||
invert H1.
|
invert H2.
|
||||||
apply inj_pair2 in H4; subst.
|
apply inj_pair2 in H5; subst.
|
||||||
cases st1'.
|
cases st1'.
|
||||||
cases p.
|
cases p.
|
||||||
eapply step_translate in H2; eauto.
|
eapply step_translate in H6; eauto.
|
||||||
first_order.
|
simp.
|
||||||
eexists; split; [ | eassumption ].
|
eexists; split; [ | eassumption ].
|
||||||
econstructor.
|
econstructor.
|
||||||
assumption.
|
assumption.
|
||||||
|
eauto.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Hint Constructors eval DE.step.
|
Hint Constructors eval DE.step.
|
||||||
|
@ -1339,12 +1355,13 @@ Module MixedToDeep(Import BW : BIT_WIDTH).
|
||||||
econstructor.
|
econstructor.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Theorem hoare_triple_sound : forall P (c : cmd wrd) Q V s,
|
Theorem hoare_triple_sound : forall P (c : cmd wrd) Q V s h H,
|
||||||
hoare_triple P c Q
|
hoare_triple P c Q
|
||||||
-> P $0
|
-> P h
|
||||||
|
-> heaps_compat H h
|
||||||
-> translate "result" V c s
|
-> translate "result" V c s
|
||||||
-> V $? "result" = None
|
-> V $? "result" = None
|
||||||
-> invariantFor (DE.trsys_of $0 V s)
|
-> invariantFor (DE.trsys_of H V s)
|
||||||
(fun p => snd p = Skip
|
(fun p => snd p = Skip
|
||||||
\/ exists p', DE.step p p').
|
\/ exists p', DE.step p p').
|
||||||
Proof.
|
Proof.
|
||||||
|
@ -1353,23 +1370,24 @@ Module MixedToDeep(Import BW : BIT_WIDTH).
|
||||||
eapply invariant_simulates.
|
eapply invariant_simulates.
|
||||||
apply translated_simulates.
|
apply translated_simulates.
|
||||||
eassumption.
|
eassumption.
|
||||||
apply invariant_multistepify with (sys := trsys_of $0 c).
|
eassumption.
|
||||||
|
apply invariant_multistepify with (sys := trsys_of h c).
|
||||||
eauto using hoare_triple_sound.
|
eauto using hoare_triple_sound.
|
||||||
invert 1; first_order.
|
invert 1; simp.
|
||||||
|
|
||||||
cases st2; simplify; subst.
|
cases st2; simplify; subst.
|
||||||
invert H4; simplify.
|
invert H6; simplify.
|
||||||
apply inj_pair2 in H9; subst.
|
apply inj_pair2 in H8; subst.
|
||||||
invert H8.
|
invert H11.
|
||||||
apply inj_pair2 in H4; subst.
|
apply inj_pair2 in H6; subst.
|
||||||
right; eexists.
|
right; eexists.
|
||||||
econstructor; eauto.
|
econstructor; eauto.
|
||||||
auto.
|
auto.
|
||||||
|
|
||||||
invert H4.
|
invert H6; simp.
|
||||||
apply inj_pair2 in H6; subst; simplify.
|
apply inj_pair2 in H8; subst; simplify.
|
||||||
cases x.
|
cases x.
|
||||||
eapply not_stuck in H3; eauto.
|
eapply not_stuck in H7; eauto.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
Theorem adder_ok : forall a b c,
|
Theorem adder_ok : forall a b c,
|
||||||
|
@ -1398,6 +1416,7 @@ 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.
|
||||||
|
|
Loading…
Reference in a new issue