@@ -798,17 +798,38 @@ Module TLB.
798798 #[export] Typeclasses Transparent Ctxt.t.
799799
800800 Module Entry.
801- Definition t (lvl : Level) := vec val (S lvl).
802- Definition pte {lvl} (tlbe : t lvl) := Vector.last tlbe.
801+ Record t (lvl : Level) :=
802+ make {
803+ val_ttbr : val;
804+ ptes : vec val (S lvl);
805+ }.
806+ Arguments make {_} _ _.
807+ Arguments val_ttbr {_}.
808+ Arguments ptes {_}.
809+
810+ #[global] Instance eq_dec lvl : EqDecision (t lvl).
811+ Proof . solve_decision. Defined .
812+
813+ #[global] Instance eqdep_dec : EqDepDecision t.
814+ Proof . intros ? ? ? [] []. decide_jmeq. Defined .
815+
816+ #[global] Instance count lvl : Countable (t lvl).
817+ Proof .
818+ eapply (inj_countable' (fun ent => (val_ttbr ent, ptes ent))
819+ (fun x => make x.1 x.2)).
820+ abstract sauto.
821+ Defined .
822+
823+ Definition pte {lvl} (tlbe : t lvl) := Vector.last tlbe.(ptes).
803824
804825 Program Definition append {lvl clvl : Level}
805826 (tlbe : t lvl)
806827 (pte : val)
807828 (CHILD : lvl + 1 = clvl) : t clvl :=
808- ctrans _ (tlbe +++ [#pte]).
829+ make tlbe.(val_ttbr) ( ctrans _ (tlbe.(ptes) +++ [#pte]) ).
809830 Solve All Obligations with lia.
810831 End Entry.
811- #[export] Typeclasses Transparent Entry.t .
832+ Export (hints) Entry.
812833
813834 (* Full Entry *)
814835 Module FE.
@@ -939,7 +960,8 @@ Module TLB.
939960 let asid := bv_extract 48 16 val_ttbr in
940961 let ndctxt := NDCtxt.make va (Some asid) in
941962 let ctxt := existT root_lvl ndctxt in
942- let entry : Entry.t (Ctxt.lvl ctxt) := [#memval] in
963+ let entry : Entry.t (Ctxt.lvl ctxt) :=
964+ Entry.make val_ttbr ([#memval] : vec val (S root_lvl)) in
943965 (* add the entry to vatlb only when it is not in the original vatlb *)
944966 if decide (entry ∉ (VATLB.get ctxt vatlb)) then
945967 Ok (VATLB.insert ctxt entry vatlb, true)
@@ -963,8 +985,8 @@ Module TLB.
963985 (te : Entry.t (Ctxt.lvl ctxt))
964986 (index : bv 9)
965987 (ttbr : reg) : result string (VATLB.t * bool) :=
966- guard_or "ASID is not active"
967- $ is_active_asid ts (Ctxt.asid ctxt) ttbr time;;
988+ (* guard_or "ASID is not active"
989+ $ is_active_asid ts (Ctxt.asid ctxt) ttbr time;; *)
968990 guard_or "The translation entry is not in the TLB"
969991 (te ∈ VATLB.get ctxt vatlb);;
970992
@@ -1192,7 +1214,7 @@ Module TLB.
11921214 else filter (λ te, is_block (TLB.Entry.pte te)) tes in
11931215 for te in (elements tes) do
11941216 ti ← invalidation_time mem tid trans_time ctxt te;
1195- mret ((vec_to_list te ), ti)
1217+ mret ((vec_to_list (Entry.ptes te) ), ti)
11961218 end .
11971219
11981220 (** Get all the TLB entries that could translate the given VA
@@ -1254,7 +1276,7 @@ Module TLB.
12541276 if decide (is_valid memval) then mret None
12551277 else
12561278 ti ← invalidation_time mem tid trans_time ctxt te;
1257- mret $ Some ((vec_to_list te ) ++ [memval], ti)
1279+ mret $ Some ((vec_to_list (Entry.ptes te) ) ++ [memval], ti)
12581280 else
12591281 mthrow "The PTE is missing"
12601282 end ;
@@ -1289,9 +1311,11 @@ Module TLB.
12891311 (va : bv 64) (asid : bv 16)
12901312 (ttbr : reg) : result string (list (list val * (option nat))) :=
12911313 candidates_asid ←
1292- get_invalid_ptes_with_inv_time_by_lvl_asid ts init mem tid tlb trans_time lvl va (Some asid) ttbr;
1314+ get_invalid_ptes_with_inv_time_by_lvl_asid
1315+ ts init mem tid tlb trans_time lvl va (Some asid) ttbr;
12931316 candidates_global ←
1294- get_invalid_ptes_with_inv_time_by_lvl_asid ts init mem tid tlb trans_time lvl va None ttbr;
1317+ get_invalid_ptes_with_inv_time_by_lvl_asid
1318+ ts init mem tid tlb trans_time lvl va None ttbr;
12951319 mret (candidates_asid ++ candidates_global).
12961320
12971321 Definition get_invalid_ptes_with_inv_time (ts : TState.t) (init : Memory.initial)
@@ -1760,8 +1784,10 @@ Definition run_trans_start (trans_start : TranslationStartInfo)
17601784 trans_res ←
17611785 if decide (va_in_range va) then
17621786 ttbr ← mlift $ ttbr_of_regime va trans_start.(TranslationStartInfo_regime);
1763- snapshots ← mlift $ TLB.unique_snapshots_until_timestamp ts init mem vmax_t va ttbr;
1764- valid_entries ← mlift $ TLB.get_valid_entries_from_snapshots snapshots mem tid va asid;
1787+ snapshots ←
1788+ mlift $ TLB.unique_snapshots_until_timestamp ts init mem vmax_t va ttbr;
1789+ valid_entries ←
1790+ mlift $ TLB.get_valid_entries_from_snapshots snapshots mem tid va asid;
17651791 invalid_entries ← mlift $
17661792 TLB.get_invalid_entries_from_snapshots snapshots ts init mem tid is_ets2 va asid ttbr;
17671793 (* update IIS with either a valid translation result or a fault *)
@@ -1836,8 +1862,8 @@ Definition run_take_exception (fault : exn) (vmax_t : view) :
18361862
18371863(** Runs an outcome. *)
18381864Section RunOutcome.
1839- Context (tid : nat) (initmem : memoryMap).
18401865
1866+ Context (tid : nat) (initmem : memoryMap).
18411867 Equations run_outcome (out : outcome) (mem_update : bool) :
18421868 Exec.t (PPState.t TState.t Ev.t IIS.t) string (eff_ret out * option view) :=
18431869 | RegRead reg racc, mem_update =>
0 commit comments