Skip to content

Commit

Permalink
Finish tulip basic resolve prophecy triples
Browse files Browse the repository at this point in the history
  • Loading branch information
jtassarotti committed Dec 8, 2024
1 parent 0eea7df commit ca87c20
Show file tree
Hide file tree
Showing 2 changed files with 35 additions and 4 deletions.
24 changes: 20 additions & 4 deletions src/program_proof/tulip/program/txn/proph.v
Original file line number Diff line number Diff line change
Expand Up @@ -49,6 +49,21 @@ Section proph.
iModIntro. by iApply "HΦ".
Qed.

Lemma decode_dbmap_map_val mv (wrs : dbmap) :
map.map_val mv = Some (Map.untype wrs) →
decode_dbmap mv = wrs.
Proof.
rewrite /decode_dbmap => ->.
rewrite /Map.untype.
rewrite -map_fmap_compose.
apply map_leibniz.
intros k.
rewrite ?lookup_fmap.
destruct (wrs !! k) as [v|] eqn:Heq; rewrite Heq //=.
rewrite /to_dbstring/dbval_to_val.
destruct v => //=.
Qed.

Lemma wp_ResolveCommit
p (tid : u64) (ts : nat) (wrsP : loc) q (wrs : dbmap) :
Expand All @@ -74,12 +89,13 @@ Section proph.
iMod ("Hclose" with "[Hp]") as "HΦ".
{ iExists (decode_actions pvs').
rewrite Hts in Hpvs.
iSplit.
{ admit. (* todo: bad definition of decode dbmap *) }
iExists _. by iFrame. }
rewrite (decode_dbmap_map_val mv wrs) // in Hpvs.
iSplit; first by naive_solver.
iExists _. by iFrame.
}
iModIntro. iApply "HΦ".
iFrame "Hmref".
eauto.
Admitted.
Qed.

End proph.
15 changes: 15 additions & 0 deletions src/program_proof/tulip/res_txnsys.v
Original file line number Diff line number Diff line change
Expand Up @@ -637,12 +637,27 @@ Section res.
Definition to_dbval (b : bool) (v : string) :=
if b then Some v else None.

Definition to_dbstring (v : val) : option string :=
match v with
| (#true, (#(LitString key), _))%V => Some key
| (#false, (#(LitString key), _))%V => None
| _ => None
end.

Definition decode_dbmap (v: val) : dbmap :=
match @map.map_val _ dbkey (@string_IntoVal grove_op) String.eq_dec String.countable v with
| None => ∅
| Some (mv, _) => to_dbstring <$> mv
end.

(*
Fixpoint decode_dbmap (v : val) : dbmap :=
match v with
| (#(LitString key), #(LitBool present), #(LitString str'), tail)%V =>
<[key:=to_dbval present str']> (decode_dbmap tail)
| _ => ∅
end.
*)

Local Definition decode_ev_read (v : val) : option action :=
match v with
Expand Down

0 comments on commit ca87c20

Please sign in to comment.