@@ -22,6 +22,7 @@ open FStar.PCM
2222module T = FStar.Tactics.V2
2323open Pulse.Lib.Dv {}
2424open FStar.ExtractAs
25+ include Pulse.Lib.Loc
2526
2627(* Arguments of slprops can be marked as a matching key to
2728 1- Make sure we do no try to use the SMT to match resources with
@@ -175,8 +176,6 @@ let inames_subset (is1 is2 : inames) : Type0 =
175176let ( /!) ( is1 is2 : inames ) : Type0 =
176177 GhostSet. disjoint is1 is2
177178
178- val inv ( i : iname ) ( p : slprop ) : slprop
179-
180179let mem_iname ( e : inames ) ( i : iname ) : erased bool = elift2 ( fun e i -> GhostSet. mem i e ) e i
181180let mem_inv ( e : inames ) ( i : iname ) : GTot bool = mem_iname e i
182181
@@ -226,11 +225,6 @@ val frame_stt
226225 ( e : stt a pre post )
227226: stt a ( pre ** frame ) ( fun x -> post x ** frame )
228227
229- val fork
230- (# pre : slprop )
231- ( f :unit -> stt unit pre ( fun _ -> emp ))
232- : stt unit pre ( fun _ -> emp )
233-
234228val sub_stt (# a :Type u# a )
235229 (# pre1 : slprop )
236230 ( pre2 : slprop )
@@ -441,12 +435,44 @@ val sub_invs_ghost
441435 ( _ : squash ( inames_subset opens1 opens2 ))
442436: stt_ghost a opens2 pre post
443437
438+ ////////////////////////////////////////////////////////////////////
439+ // Locations
440+ ////////////////////////////////////////////////////////////////////
441+
442+ val loc : loc_id -> timeless_slprop
443+
444+ val loc_get () : stt_ghost loc_id emp_inames emp ( fun l -> loc l )
445+ val loc_dup l : stt_ghost unit emp_inames ( loc l ) ( fun _ -> loc l ** loc l )
446+ val loc_gather l # l' : stt_ghost unit emp_inames ( loc l ** loc l' ) ( fun _ -> loc l ** pure ( l == l' ))
447+
448+ val on ( l : loc_id ) ([ @@@mkey ] p : slprop ) : slprop
449+ val on_intro # l p : stt_ghost unit emp_inames ( loc l ** p ) ( fun _ -> loc l ** on l p )
450+ val on_elim # l p : stt_ghost unit emp_inames ( loc l ** on l p ) ( fun _ -> loc l ** p )
451+
452+ val timeless_on ( l : loc_id ) ( p : slprop )
453+ : Lemma
454+ ( requires timeless p )
455+ ( ensures timeless ( on l p ))
456+ [ SMTPat ( timeless ( on l p ))]
457+
458+ val on_star_eq l a b : squash ( on l ( a ** b ) == on l a ** on l b )
459+ val on_on_eq l1 l2 a : squash ( on l1 ( on l2 a ) == on l2 a )
460+ val on_loc_eq l1 l2 : squash ( on l1 ( loc l2 ) == pure ( l1 == l2 ))
461+
462+ val ghost_impersonate_core
463+ (#[ T. exact (` emp_inames )] is : inames )
464+ ( l : loc_id ) ( pre post : slprop )
465+ ( f : unit -> stt_ghost unit is pre ( fun _ -> post ))
466+ : stt_ghost unit is ( on l pre ) ( fun _ -> on l post )
467+
444468//////////////////////////////////////////////////////////////////////////
445469// Later
446470//////////////////////////////////////////////////////////////////////////
447471
448472val later_credit ( amt : nat) : slprop
449473
474+ val on_later_credit_eq l n : squash ( on l ( later_credit n ) == later_credit n )
475+
450476val timeless_later_credit ( amt : nat)
451477: Lemma ( timeless ( later_credit amt ))
452478 [ SMTPat ( timeless ( later_credit amt ))]
@@ -471,13 +497,17 @@ val later_star p q : squash (later (p ** q) == later p ** later q)
471497val later_exists (# t : Type) ( f : t -> slprop ) : stt_ghost unit emp_inames ( later ( exists * x . f x )) ( fun _ -> exists * x . later ( f x ))
472498val exists_later (# t : Type) ( f : t -> slprop ) : stt_ghost unit emp_inames ( exists * x . later ( f x )) ( fun _ -> later ( exists * x . f x ))
473499
500+ val on_later_eq l p : squash ( on l ( later p ) == later ( on l p ))
501+
474502//////////////////////////////////////////////////////////////////////////
475503// Equivalence
476504//////////////////////////////////////////////////////////////////////////
477505
478506(* Two slprops are equal when approximated to the current heap level. *)
479507val equiv ( a b : slprop ) : slprop
480508
509+ val on_equiv_eq l a b : squash ( on l ( equiv a b ) == equiv a b )
510+
481511val equiv_dup a b : stt_ghost unit emp_inames ( equiv a b ) fun _ -> equiv a b ** equiv a b
482512val equiv_refl a : stt_ghost unit emp_inames emp fun _ -> equiv a a
483513val equiv_comm a b : stt_ghost unit emp_inames ( equiv a b ) fun _ -> equiv b a
@@ -502,6 +532,8 @@ val null_slprop_ref : slprop_ref
502532
503533val slprop_ref_pts_to ([ @@@mkey ] x : slprop_ref ) ( y : slprop ) : slprop
504534
535+ val on_slprop_ref_pts_to_eq l x y : squash ( on l ( slprop_ref_pts_to x y ) == slprop_ref_pts_to x y )
536+
505537val slprop_ref_alloc ( y : slprop )
506538: stt_ghost slprop_ref emp_inames emp fun x -> slprop_ref_pts_to x y
507539
@@ -512,57 +544,6 @@ val slprop_ref_share (x: slprop_ref) (#y: slprop)
512544val slprop_ref_gather ( x : slprop_ref ) (# y1 # y2 : slprop )
513545: stt_ghost unit emp_inames ( slprop_ref_pts_to x y1 ** slprop_ref_pts_to x y2 ) fun _ -> slprop_ref_pts_to x y1 ** later ( equiv y1 y2 )
514546
515- //////////////////////////////////////////////////////////////////////////
516- // Invariants
517- //////////////////////////////////////////////////////////////////////////
518-
519- val dup_inv ( i : iname ) ( p : slprop )
520- : stt_ghost unit emp_inames ( inv i p ) ( fun _ -> inv i p ** inv i p )
521-
522- val new_invariant ( p : slprop )
523- : stt_ghost iname emp_inames p ( fun i -> inv i p )
524-
525- val fresh_invariant
526- ( ctx : inames { Pulse.Lib.GhostSet. is_finite ctx })
527- ( p : slprop )
528- : stt_ghost ( i : iname { ~( i ` GhostSet. mem ` ctx ) }) emp_inames p ( fun i -> inv i p )
529-
530- val with_invariant
531- (# a :Type)
532- (# obs : _ )
533- (# fp : slprop )
534- (# fp' : a -> slprop )
535- (# f_opens : inames )
536- (# p : slprop )
537- ( i : iname { not ( mem_inv f_opens i ) })
538- ($ f :unit -> stt_atomic a # obs f_opens
539- ( later p ** fp )
540- ( fun x -> later p ** fp' x ))
541- : stt_atomic a # obs ( add_inv f_opens i ) ( inv i p ** fp ) ( fun x -> inv i p ** fp' x )
542-
543- val with_invariant_g
544- (# a :Type)
545- (# fp : slprop )
546- (# fp' : a -> slprop )
547- (# f_opens : inames )
548- (# p : slprop )
549- ( i : iname { not ( mem_inv f_opens i ) })
550- ($ f :unit -> stt_ghost a f_opens
551- ( later p ** fp )
552- ( fun x -> later p ** fp' x ))
553- : stt_ghost a ( add_inv f_opens i ) ( inv i p ** fp ) ( fun x -> inv i p ** fp' x )
554-
555- [ @@allow_ambiguous ]
556- val invariant_name_identifies_invariant
557- (# p # q : slprop )
558- ( i : iname )
559- ( j : iname { i == j } )
560- : stt_ghost
561- unit
562- emp_inames
563- ( inv i p ** inv j q )
564- ( fun _ -> inv i p ** inv j q ** later ( equiv p q ))
565-
566547(* **** end computation types and combinators *****)
567548
568549(* This tactic is called to find non_informative witnesses.
@@ -575,6 +556,11 @@ let non_info_tac () : T.Tac unit =
575556// Some basic actions and ghost operations
576557//////////////////////////////////////////////////////////////////////////
577558
559+ val fork_core
560+ ( pre : slprop ) # l
561+ ( f : ( l' : loc_id { process_of l' == process_of l } -> stt unit ( loc l' ** on l pre ) ( fun _ -> emp )))
562+ : stt unit ( loc l ** pre ) ( fun _ -> emp )
563+
578564val rewrite ( p : slprop ) ( q : slprop ) ( _ : slprop_equiv p q )
579565: stt_ghost unit emp_inames p ( fun _ -> q )
580566
@@ -670,7 +656,8 @@ val elim_false (a:Type) (p:a -> slprop)
670656
671657// Finally, a big escape hatch for introducing architecture/backend-specific
672658// atomic operations from proven stt specifications
673- [ @@warn_on_use " as_atomic is a an assumption" ]
659+ [ @@warn_on_use " as_atomic is a an assumption" ;
660+ extract_as (`( fun (# a : Type0) ( pre post : unit) ( f : a ) -> f ))]
674661val as_atomic (# a :Type u# 0 ) ( pre : slprop ) ( post : a -> slprop )
675662 ( pf : stt a pre post )
676663: stt_atomic a emp_inames pre post
0 commit comments