Skip to content

Commit 53d1c39

Browse files
committed
Simplify promise-first code
- Apply code reviews Apply code reviews
1 parent fbfa855 commit 53d1c39

File tree

5 files changed

+108
-345
lines changed

5 files changed

+108
-345
lines changed

ArchSem/GenPromising.v

Lines changed: 27 additions & 53 deletions
Original file line numberDiff line numberDiff line change
@@ -380,7 +380,7 @@ Module GenPromising (IWA : InterfaceWithArch) (TM : TermModelsT IWA).
380380
promise ← mchoosel (enum bool);
381381
if (promise : bool) then cpromise_tid fuel tid else run_tid isem prom tid.
382382

383-
(** Computational evaluate all the possible allowed final states according
383+
(** Computationally evaluate all the possible allowed final states according
384384
to the promising model prom starting from st *)
385385
Fixpoint run (fuel : nat) : Exec.t t string final :=
386386
st ← mGet;
@@ -391,54 +391,41 @@ Module GenPromising (IWA : InterfaceWithArch) (TM : TermModelsT IWA).
391391
run fuel
392392
else mthrow "Could not finish running within the size of the fuel".
393393

394-
(** Explore all possible promise-based executions across all threads. *)
395-
Fixpoint prune_promises_and_states (fuel_per_tid fuel : nat)
396-
(finals : list final) : Exec.t t string (list final) :=
394+
(** Computationally evaluate all the possible allowed final states according
395+
to the promising model prom starting from st with promise-first optimization.
396+
The size of fuel should be at least (# of promises) + max(# of instructions) + 1 *)
397+
Fixpoint run_promise_first (fuel : nat) : Exec.t t string final :=
397398
if fuel is S fuel then
398399
st ← mGet;
399400
(* Find out next possible promises or terminating states at the current thread *)
400401
executions ←
401-
for (tid : fin n) in enum (fin n) do
402+
vmapM (λ '(tid, ts),
402403
'(promises_per_tid, tstates_per_tid) ←
403404
mlift $ prom.(enumerate_promises_and_terminal_states)
404-
fuel_per_tid tid (term tid) (initmem st) (tstate tid st) (events st);
405-
mret (map (λ ev, (ev, tid)) promises_per_tid, tstates_per_tid)
406-
end;
407-
408-
(* Compute cartesian products of the possible thread states *)
409-
let tstates_sys :=
410-
fold_left (λ partial_sts tstates_tid,
411-
List.flat_map (λ tstate,
412-
if is_emptyb partial_sts then [[tstate]]
413-
else map (λ partial_st, partial_st ++ [tstate]) partial_sts
414-
) tstates_tid
415-
) (map snd executions) [] in
416-
let new_finals :=
417-
omap (λ tstates,
418-
if (list_to_vec_n n tstates) is Some tstates_vec then
419-
let st := Make tstates_vec st.(PState.initmem) st.(PState.events) in
420-
if decide $ terminated prom term st is left pt
421-
then Some (make_final st pt)
422-
else None
423-
else None
424-
) tstates_sys in
405+
fuel (fin_to_nat tid) (term tid) (initmem st) ts (events st);
406+
mret (map (.,tid) promises_per_tid, tstates_per_tid)
407+
) (venumerate (tstates st));
425408

426-
(* Non-deterministically choose the next promise and the tid for pruning *)
427-
let promises_all := List.concat (map fst executions) in
428-
if is_emptyb promises_all then
429-
mret (new_finals ++ finals)
430-
else
409+
promise ← mchoosel (enum bool);
410+
if (promise : bool) then
411+
(* Non-deterministically choose the next promise and the tid for pruning *)
412+
let promises_all := List.concat (vmap fst executions) in
431413
'(next_ev, tid) ← mchoosel promises_all;
432414
mSet (λ st, promise_tid prom st tid next_ev);;
433-
prune_promises_and_states fuel_per_tid fuel (new_finals ++ finals)
434-
else mret finals.
435-
436-
(** Computational evaluate all the possible allowed final states according
437-
to the promising model prom starting from st
438-
with promise-first optimization *)
439-
Definition run_promise_first (fuel : nat) : Exec.t t string final :=
440-
finals ← prune_promises_and_states fuel fuel [];
441-
mchoosel finals.
415+
run_promise_first fuel
416+
else
417+
(* Compute cartesian products of the possible thread states *)
418+
let tstates_all := cprodn (vmap snd executions) in
419+
let new_finals :=
420+
omap (λ tstates,
421+
let st := Make tstates st.(PState.initmem) st.(PState.events) in
422+
if decide $ terminated prom term st is left pt
423+
then Some (make_final st pt)
424+
else None
425+
) tstates_all in
426+
mchoosel new_finals
427+
else
428+
mthrow "Could not finish running within the size of the fuel".
442429

443430
End CPS.
444431
Arguments to_final_archState {_ _ _}.
@@ -447,25 +434,12 @@ Module GenPromising (IWA : InterfaceWithArch) (TM : TermModelsT IWA).
447434
(** Create a computational model from an ISA model and promising model *)
448435
Definition Promising_to_Modelc (isem : iMon ()) (prom : BasicExecutablePM)
449436
(fuel : nat) : archModel.c ∅ :=
450-
λ n term (initMs : archState n),
451-
PState.from_archState prom initMs |>
452-
archModel.Res.from_exec
453-
$ CPState.to_final_archState
454-
<$> CPState.run isem prom term fuel.
455-
456-
(** Create a computational model from an ISA model and promising model
457-
with promise-free optimization *)
458-
Definition Promising_to_Modelc_pf (isem : iMon ()) (prom : BasicExecutablePM)
459-
(fuel : nat) : archModel.c ∅ :=
460437
λ n term (initMs : archState n),
461438
PState.from_archState prom initMs |>
462439
archModel.Res.from_exec
463440
$ CPState.to_final_archState
464441
<$> CPState.run_promise_first prom term fuel.
465442

466-
(* TODO state some soundness lemma between Promising_to_Modelnc and
467-
Promising_Modelc *)
468-
469443
End GenPromising.
470444

471445
Module Type GenPromisingT (IWA : InterfaceWithArch) (TM : TermModelsT IWA).

ArchSemArm/UMPromising.v

Lines changed: 35 additions & 86 deletions
Original file line numberDiff line numberDiff line change
@@ -457,18 +457,16 @@ Definition read_mem4 (addr : address) (macc : mem_acc) (init : Memory.initial) :
457457
This may mutate memory if no existing promise can be fullfilled *)
458458
Definition write_mem (tid : nat) (loc : Loc.t) (vdata : view)
459459
(macc : mem_acc) (mem : Memory.t)
460-
(data : val) (mem_update : bool) :
460+
(data : val) :
461461
Exec.t TState.t string (Memory.t * view * option view):=
462462
let msg := Msg.make tid loc data in
463463
let is_release := is_rel_acq macc in
464464
ts ← mGet;
465-
'(time, mem, new_promise)
465+
let '(time, mem, new_promise) :=
466466
match Memory.fulfill msg (TState.prom ts) mem with
467-
| Some t => mret (t, mem, false)
468-
| None =>
469-
if mem_update then mret (Memory.promise msg mem, true)
470-
else mdiscard
471-
end;
467+
| Some t => (t, mem, false)
468+
| None => (Memory.promise msg mem, true)
469+
end in
472470
let vbob :=
473471
ts.(TState.vdmbst) ⊔ ts.(TState.vdmb) ⊔ ts.(TState.visb) ⊔ ts.(TState.vacq)
474472
⊔ view_if is_release (ts.(TState.vrd) ⊔ ts.(TState.vwr)) in
@@ -478,10 +476,8 @@ Definition write_mem (tid : nat) (loc : Loc.t) (vdata : view)
478476
mSet $ TState.update_coh loc time;;
479477
mSet $ TState.update TState.vwr time;;
480478
mSet $ TState.update TState.vrel (view_if is_release time);;
481-
mret $ match new_promise with
482-
| true => (mem, time, Some vpre)
483-
| false => (mem, time, None)
484-
end.
479+
mret (mem, time, (if new_promise then Some vpre else None)).
480+
485481

486482
(** Tries to perform a memory write.
487483
@@ -492,12 +488,12 @@ Definition write_mem (tid : nat) (loc : Loc.t) (vdata : view)
492488
return value indicate the success (true for success, false for error) *)
493489
Definition write_mem_xcl (tid : nat) (loc : Loc.t)
494490
(vdata : view) (macc : mem_acc)
495-
(mem : Memory.t) (data : val) (mem_update : bool)
491+
(mem : Memory.t) (data : val)
496492
: Exec.t TState.t string (Memory.t * option view) :=
497493
guard_or "Atomic RMW unsupported" (¬ (is_atomic_rmw macc));;
498494
let xcl := is_exclusive macc in
499495
if xcl then
500-
'(mem, time, vpre_opt) ← write_mem tid loc vdata macc mem data mem_update;
496+
'(mem, time, vpre_opt) ← write_mem tid loc vdata macc mem data;
501497
ts ← mGet;
502498
match TState.xclb ts with
503499
| None => mdiscard
@@ -508,7 +504,7 @@ Definition write_mem_xcl (tid : nat) (loc : Loc.t)
508504
mSet TState.clear_xclb;;
509505
mret (mem, vpre_opt)
510506
else
511-
'(mem, time, vpre_opt) ← write_mem tid loc vdata macc mem data mem_update;
507+
'(mem, time, vpre_opt) ← write_mem tid loc vdata macc mem data;
512508
mSet $ TState.set_fwdb loc (FwdItem.make time vdata false);;
513509
mret (mem, vpre_opt).
514510

@@ -538,9 +534,9 @@ End IIS.
538534
Section RunOutcome.
539535
Context (tid : nat) (initmem : memoryMap).
540536

541-
Equations run_outcome (out : outcome) (mem_update : bool) :
537+
Equations run_outcome (out : outcome) :
542538
Exec.t (PPState.t TState.t Msg.t IIS.t) string (eff_ret out * option view) :=
543-
| RegWrite reg racc val, mem_update =>
539+
| RegWrite reg racc val =>
544540
guard_or "Non trivial reg access types unsupported" (racc = None);;
545541
vreg ← mget (IIS.strict ∘ PPState.iis);
546542
vreg' ←
@@ -554,14 +550,14 @@ Section RunOutcome.
554550
TState.set_reg reg (val, vreg') ts;
555551
msetv PPState.state nts;;
556552
mret ((), None)
557-
| RegRead reg racc, mem_update =>
553+
| RegRead reg racc =>
558554
guard_or "Non trivial reg access types unsupported" (racc = None);;
559555
ts ← mget PPState.state;
560556
'(val, view) ← othrow "Register isn't mapped can't read" $
561557
dmap_lookup reg ts.(TState.regs);
562558
mset PPState.iis $ IIS.add view;;
563559
mret (val, None)
564-
| MemRead (MemReq.make macc addr addr_space 8 0), mem_update =>
560+
| MemRead (MemReq.make macc addr addr_space 8 0) =>
565561
guard_or "Access outside Non-Secure" (addr_space = PAS_NonSecure);;
566562
loc ← othrow "PA not supported" $ Loc.from_addr addr;
567563
if is_ifetch macc then
@@ -575,28 +571,28 @@ Section RunOutcome.
575571
mset PPState.iis $ IIS.add view;;
576572
mret (Ok (val, bv_0 0), None)
577573
else mthrow "Read is not explicit or ifetch"
578-
| MemRead (MemReq.make macc addr addr_space 4 0), mem_update => (* ifetch *)
574+
| MemRead (MemReq.make macc addr addr_space 4 0) => (* ifetch *)
579575
guard_or "Access outside Non-Secure" (addr_space = PAS_NonSecure);;
580576
let initmem := Memory.initial_from_memMap initmem in
581577
opcode ← Exec.liftSt PPState.mem $ read_mem4 addr macc initmem;
582578
mret (Ok (opcode, 0%bv), None)
583-
| MemRead _, mem_update => mthrow "Memory read of size other than 8 and 4"
584-
| MemWriteAddrAnnounce _, mem_update =>
579+
| MemRead _ => mthrow "Memory read of size other than 8 and 4"
580+
| MemWriteAddrAnnounce _ =>
585581
vaddr ← mget (IIS.strict ∘ PPState.iis);
586582
mset PPState.state $ TState.update TState.vcap vaddr;;
587583
mret ((), None)
588-
| MemWrite (MemReq.make macc addr addr_space 8 0) val tags, mem_update =>
584+
| MemWrite (MemReq.make macc addr addr_space 8 0) val tags =>
589585
guard_or "Access outside Non-Secure" (addr_space = PAS_NonSecure);;
590586
addr ← othrow "PA not supported" $ Loc.from_addr addr;
591587
if is_explicit macc then
592588
mem ← mget PPState.mem;
593589
vdata ← mget (IIS.strict ∘ PPState.iis);
594590
'(mem, vpre_opt) ← Exec.liftSt PPState.state
595-
$ write_mem_xcl tid addr vdata macc mem val mem_update;
591+
$ write_mem_xcl tid addr vdata macc mem val;
596592
msetv PPState.mem mem;;
597593
mret (Ok (), vpre_opt)
598594
else mthrow "Unsupported non-explicit write"
599-
| Barrier (Barrier_DMB dmb), mem_update => (* dmb *)
595+
| Barrier (Barrier_DMB dmb) => (* dmb *)
600596
ts ← mget PPState.state;
601597
match dmb.(DxB_types) with
602598
| MBReqTypes_All (* dmb sy *) =>
@@ -607,16 +603,17 @@ Section RunOutcome.
607603
mset PPState.state $ TState.update TState.vdmbst ts.(TState.vwr)
608604
end;;
609605
mret ((), None)
610-
| Barrier (Barrier_ISB ()), mem_update => (* isb *)
606+
| Barrier (Barrier_ISB ()) => (* isb *)
611607
ts ← mget PPState.state;
612608
mset PPState.state $ TState.update TState.visb (TState.vcap ts);;
613609
mret ((), None)
614-
| GenericFail s, mem_update => mthrow ("Instruction failure: " ++ s)%string
615-
| _, _ => mthrow "Unsupported outcome".
610+
| GenericFail s => mthrow ("Instruction failure: " ++ s)%string
611+
| _ => mthrow "Unsupported outcome".
616612

617613
Definition run_outcome' (out : outcome) :
618614
Exec.t (PPState.t TState.t Msg.t IIS.t) string (eff_ret out) :=
619-
run_outcome out true |$> fst.
615+
run_outcome out |$> fst.
616+
620617
End RunOutcome.
621618

622619
Module CProm.
@@ -655,14 +652,17 @@ Section ComputeProm.
655652
(base : view)
656653
(out : outcome) :
657654
Exec.t (CProm.t * PPState.t TState.t Msg.t IIS.t) string (eff_ret out) :=
658-
'(res, vpre_opt) ← Exec.liftSt snd $ run_outcome tid initmem out true;
655+
'(res, vpre_opt) ← Exec.liftSt snd $ run_outcome tid initmem out;
659656
if vpre_opt is Some vpre then
660657
mem ← mget (PPState.mem ∘ snd);
661658
mset fst (CProm.add_if mem vpre base);;
662659
mret res
663660
else
664661
mret res.
665662

663+
(* Run the thread state until termination, collecting certified promises.
664+
Returns true if termination occurs within the given fuel,
665+
false otherwise. *)
666666
Fixpoint run_to_termination_promise
667667
(isem : iMon ())
668668
(fuel : nat)
@@ -682,43 +682,7 @@ Section ComputeProm.
682682
run_to_termination_promise isem fuel base
683683
end.
684684

685-
Definition run_to_termination (isem : iMon ())
686-
(fuel : nat)
687-
(ts : TState.t)
688-
(mem : Memory.t) :
689-
result string (list Msg.t * list TState.t) :=
690-
let base := List.length mem in
691-
let res := Exec.results $
692-
run_to_termination_promise isem fuel base (CProm.init, PPState.Make ts mem IIS.init) in
693-
guard_or ("Could not finish promises within the size of the fuel")%string
694-
(∀ r ∈ res, r.2 = true);;
695-
mret $ (CProm.proms (union_list res.*1.*1), []).
696-
697-
Definition run_outcome_with_no_promise
698-
(out : outcome) :
699-
Exec.t (PPState.t TState.t Msg.t IIS.t) string (eff_ret out) :=
700-
'(res, _) ← run_outcome tid initmem out false;
701-
mret res.
702-
703-
Fixpoint run_to_termination_no_promise
704-
(isem : iMon ())
705-
(fuel : nat) :
706-
Exec.t (PPState.t TState.t Msg.t IIS.t) string bool :=
707-
match fuel with
708-
| 0%nat =>
709-
ts ← mget PPState.state;
710-
mret (term (TState.reg_map ts))
711-
| S fuel =>
712-
let handler := run_outcome_with_no_promise in
713-
cinterp handler isem;;
714-
ts ← mget PPState.state;
715-
if term (TState.reg_map ts) then
716-
mret true
717-
else
718-
run_to_termination_no_promise isem fuel
719-
end.
720-
721-
Definition run_to_termination_pf (isem : iMon ())
685+
Definition run_to_termination (isem : iMon ())
722686
(fuel : nat)
723687
(ts : TState.t)
724688
(mem : Memory.t)
@@ -729,11 +693,11 @@ Section ComputeProm.
729693
guard_or ("Could not finish promises within the size of the fuel")%string
730694
(∀ r ∈ res_proms, r.2 = true);;
731695
let promises := res_proms.*1.*1 |> CProm.proms ∘ union_list in
732-
let res := Exec.results $
733-
run_to_termination_no_promise isem fuel (PPState.Make ts mem IIS.init) in
734-
guard_or ("Could not finish the remaining states within the size of the fuel")%string
735-
(∀ r ∈ res, r.2 = true);;
736-
let tstates := map PPState.state res.*1 in
696+
let tstates :=
697+
res_proms
698+
|> omap (λ '((cp, st), _),
699+
if is_emptyb (CProm.proms cp) then Some (PPState.state st)
700+
else None) in
737701
mret (promises, tstates).
738702

739703
End ComputeProm.
@@ -819,20 +783,5 @@ Next Obligation. Admitted.
819783
Next Obligation. Admitted.
820784
Next Obligation. Admitted.
821785

822-
Program Definition UMPromising_exe_pf' (isem : iMon ())
823-
: BasicExecutablePM :=
824-
{|pModel := UMPromising_cert' isem;
825-
enumerate_promises_and_terminal_states :=
826-
λ fuel tid term initmem ts mem,
827-
run_to_termination_pf tid initmem term isem fuel ts mem
828-
|}.
829-
Next Obligation. Admitted.
830-
Next Obligation. Admitted.
831-
Next Obligation. Admitted.
832-
Next Obligation. Admitted.
833-
834786
Definition UMPromising_cert_c isem fuel :=
835787
Promising_to_Modelc isem (UMPromising_exe' isem) fuel.
836-
837-
Definition UMPromising_cert_c_pf isem fuel :=
838-
Promising_to_Modelc_pf isem (UMPromising_exe_pf' isem) fuel.

0 commit comments

Comments
 (0)