@@ -811,10 +811,12 @@ Module TLB.
811811 Module Entry.
812812 Record t (lvl : Level) :=
813813 make {
814+ is_upper : bool;
814815 val_ttbr : val;
815816 ptes : vec val (S lvl);
816817 }.
817- Arguments make {_} _ _.
818+ Arguments make {_} _ _ _.
819+ Arguments is_upper {_}.
818820 Arguments val_ttbr {_}.
819821 Arguments ptes {_}.
820822
@@ -826,8 +828,8 @@ Module TLB.
826828
827829 #[global] Instance count lvl : Countable (t lvl).
828830 Proof .
829- eapply (inj_countable' (fun ent => (val_ttbr ent, ptes ent))
830- (fun x => make x.1 x.2)).
831+ eapply (inj_countable' (fun ent => (is_upper ent, val_ttbr ent, ptes ent))
832+ (fun x => make x.1.1 x.1.2 x.2)).
831833 abstract sauto.
832834 Defined .
833835
@@ -837,7 +839,7 @@ Module TLB.
837839 (tlbe : t lvl)
838840 (pte : val)
839841 (CHILD : lvl + 1 = clvl) : t clvl :=
840- make tlbe.(val_ttbr) (ctrans _ (tlbe.(ptes) +++ [#pte])).
842+ make tlbe.(is_upper) tlbe.( val_ttbr) (ctrans _ (tlbe.(ptes) +++ [#pte])).
841843 Solve All Obligations with lia.
842844 End Entry.
843845 Export (hints) Entry.
@@ -944,6 +946,15 @@ Module TLB.
944946 (CHILD : (Ctxt.lvl ctxt) + 1 = clvl) : prefix clvl :=
945947 bv_concat (level_length clvl) (Ctxt.va ctxt) index.
946948
949+ Definition is_upper_ttbr (ttbr : reg) : option bool :=
950+ if decide
951+ (ttbr = GReg TTBR0_EL1 ∨
952+ ttbr = GReg TTBR0_EL2 ∨
953+ ttbr = GReg TTBR0_EL3) then Some false
954+ else if decide
955+ (ttbr = GReg TTBR1_EL1 ∨
956+ ttbr = GReg TTBR1_EL2) then Some true
957+ else None.
947958
948959 (** Seed root-level TLB entries from [ttbr].
949960 - Reads [ttbr] at [time], checks it is 64-bit.
@@ -959,6 +970,7 @@ Module TLB.
959970 (ttbr : reg) : result string (VATLB.t * bool) :=
960971 sregs ← othrow "TTBR should exist in initial state"
961972 $ TState.read_sreg_at ts ttbr time;
973+ is_upper ← othrow "The register is not TTBR" (is_upper_ttbr ttbr);
962974 fold_left (λ prev sreg,
963975 '(vatlb, is_changed) ← prev;
964976 val_ttbr ← othrow "TTBR should be a 64 bit value"
@@ -971,7 +983,7 @@ Module TLB.
971983 let ndctxt := NDCtxt.make va (Some asid) in
972984 let ctxt := existT root_lvl ndctxt in
973985 let entry : Entry.t (Ctxt.lvl ctxt) :=
974- Entry.make val_ttbr ([#memval] : vec val (S root_lvl)) in
986+ Entry.make is_upper val_ttbr ([#memval] : vec val (S root_lvl)) in
975987 (* add the entry to vatlb only when it is not in the original vatlb *)
976988 if decide (entry ∉ (VATLB.get ctxt vatlb)) then
977989 Ok (VATLB.add ctxt entry vatlb, true)
@@ -1216,26 +1228,27 @@ Module TLB.
12161228 (tlb : t) (trans_time : nat)
12171229 (lvl : Level)
12181230 (ndctxt : NDCtxt.t lvl) :
1219- result string (list (list val * (option nat))) :=
1231+ result string (list (val * list val * (option nat))) :=
12201232 let ctxt := existT lvl ndctxt in
12211233 let tes := VATLB.get ctxt tlb.(TLB.vatlb) in
12221234 let tes := if decide (lvl = leaf_lvl) then tes
12231235 else filter (λ te, is_block (TLB.Entry.pte te)) tes in
12241236 for te in (elements tes) do
12251237 ti ← invalidation_time mem tid trans_time ctxt te;
1226- mret ((vec_to_list (Entry.ptes te)), ti)
1238+ mret ((Entry.val_ttbr te), ( vec_to_list (Entry.ptes te)), ti)
12271239 end .
12281240
1229- (** Get all the TLB entries that could translate the given VA
1230- at the provided level and in the provided ASID context.
1241+ (** Get all the TLB entries and the corresponding TTBR value that
1242+ could translate the given VA at the provided level
1243+ and in the provided ASID context.
12311244 Return each TLB entry as a list of descriptors [list val] with
12321245 the invalidation time [nat] *)
12331246 Definition get_leaf_ptes_with_inv_time (mem : Memory.t)
12341247 (tid : nat)
12351248 (tlb : t) (trans_time : nat)
12361249 (lvl : Level)
12371250 (va : bv 64) (asid : bv 16) :
1238- result string (list (list val * (option nat))) :=
1251+ result string (list (val * list val * (option nat))) :=
12391252 let ndctxt_asid := NDCtxt.make (level_prefix va lvl) (Some asid) in
12401253 let ndctxt_global := NDCtxt.make (level_prefix va lvl) None in
12411254 candidates_asid ←
@@ -1244,21 +1257,22 @@ Module TLB.
12441257 get_leaf_ptes_with_inv_time_by_ctxt mem tid tlb trans_time lvl ndctxt_global;
12451258 mret (candidates_asid ++ candidates_global).
12461259
1247- (** Get all the TLB entries that could translate the given VA
1248- in the provided ASID context.
1260+ (** Get all the TLB entries and the corresponding TTBR value that
1261+ could translate the given VA in the provided ASID context.
12491262 Return each TLB entry as a list of descriptors [list val] with
12501263 the invalidation time [nat] *)
12511264 Definition lookup (mem : Memory.t)
12521265 (tid : nat)
12531266 (tlb : TLB.t) (trans_time : nat)
12541267 (va : bv 64) (asid : bv 16) :
1255- result string (list (list val * (option nat))) :=
1268+ result string (list (val * list val * (option nat))) :=
12561269 res1 ← get_leaf_ptes_with_inv_time mem tid tlb trans_time 1%fin va asid;
12571270 res2 ← get_leaf_ptes_with_inv_time mem tid tlb trans_time 2%fin va asid;
12581271 res3 ← get_leaf_ptes_with_inv_time mem tid tlb trans_time leaf_lvl va asid;
12591272 mret (res1 ++ res2 ++ res3).
12601273
1261- (** Get all the TLB entries that trigger fault from the given VA
1274+ (** Get all the TLB entries and the corresponding TTBR value that
1275+ trigger fault from the given VA
12621276 at the provided level and in the provided ASID context.
12631277 Return each TLB entry as a list of descriptors [list val] with
12641278 the invalidation time [option nat] *)
@@ -1270,7 +1284,8 @@ Module TLB.
12701284 (lvl : Level)
12711285 (va : bv 64)
12721286 (asid : option (bv 16))
1273- (ttbr : reg) : result string (list (list val * (option nat))) :=
1287+ (ttbr : reg) :
1288+ result string (list (val * list val * (option nat))) :=
12741289 if parent_lvl lvl is Some parent_lvl then
12751290 let ndctxt := NDCtxt.make (level_prefix va parent_lvl) asid in
12761291 let ctxt := existT parent_lvl ndctxt in
@@ -1285,7 +1300,8 @@ Module TLB.
12851300 if decide (is_valid memval) then mret None
12861301 else
12871302 ti ← invalidation_time mem tid trans_time ctxt te;
1288- mret $ Some ((vec_to_list (Entry.ptes te)) ++ [memval], ti)
1303+ let vals := (vec_to_list (Entry.ptes te)) ++ [memval] in
1304+ mret $ Some (Entry.val_ttbr te, vals, ti)
12891305 else
12901306 mthrow "The PTE is missing"
12911307 end ;
@@ -1303,22 +1319,25 @@ Module TLB.
13031319 if (Memory.read_at loc init mem trans_time) is Some (memval, _) then
13041320 if decide (is_valid memval) then mret None
13051321 else
1306- mret $ Some ([memval], None)
1322+ mret $ Some (val_ttbr, [memval], None)
13071323 else mthrow "The root PTE is missing"
13081324 end ;
13091325 mret $ omap id invalid_ptes.
13101326
1311- (** Get all the TLB entries that trigger fault from the given VA
1327+ (** Get all the TLB entries and the corresponding TTBR value that
1328+ trigger fault from the given VA
13121329 in the provided ASID context.
13131330 Return each TLB entry as a list of descriptors [list val] with
13141331 the invalidation time [option nat] *)
1315- Definition get_invalid_ptes_with_inv_time_by_lvl (ts : TState.t) (init : Memory.initial)
1332+ Definition get_invalid_ptes_with_inv_time_by_lvl (ts : TState.t)
1333+ (init : Memory.initial)
13161334 (mem : Memory.t)
13171335 (tid : nat)
13181336 (tlb : t) (trans_time : nat)
13191337 (lvl : Level)
13201338 (va : bv 64) (asid : bv 16)
1321- (ttbr : reg) : result string (list (list val * (option nat))) :=
1339+ (ttbr : reg) :
1340+ result string (list (val * list val * (option nat))) :=
13221341 candidates_asid ←
13231342 get_invalid_ptes_with_inv_time_by_lvl_asid
13241343 ts init mem tid tlb trans_time lvl va (Some asid) ttbr;
@@ -1333,7 +1352,7 @@ Module TLB.
13331352 (tlb : TLB.t) (time : nat)
13341353 (va : bv 64) (asid : bv 16)
13351354 (ttbr : reg) :
1336- result string (list (list val * (option nat))) :=
1355+ result string (list (val * list val * (option nat))) :=
13371356 fault_ptes ←
13381357 for lvl in enum Level do
13391358 get_invalid_ptes_with_inv_time_by_lvl ts init mem tid tlb time lvl va asid ttbr
@@ -1355,7 +1374,8 @@ Module IIS.
13551374 make {
13561375 va : bv 36;
13571376 time : nat;
1358- remaining : list (bv 64); (* NOTE: translation memory read - ptes *)
1377+ root : option {ttbr : reg & reg_type ttbr};
1378+ remaining : list (bv 64);
13591379 invalidation : option nat
13601380 }.
13611381
@@ -1434,26 +1454,50 @@ Definition read_pte (vaddr : view) :
14341454 mset fst $ TState.update TState.vspec vpost;;
14351455 mret (vpost, val).
14361456
1457+ Definition run_reg_general_read (reg : reg) (racc : reg_acc) :
1458+ Exec.t (TState.t * IIS.t) string (reg_type reg * view) :=
1459+ ts ← mget fst;
1460+ if decide (reg ∈ relaxed_regs) then
1461+ if decide (is_Some racc)
1462+ then othrow
1463+ ("Register " ++ pretty reg ++ " unmapped on direct read")%string
1464+ $ TState.read_sreg_direct ts reg
1465+ else
1466+ valvs ← othrow
1467+ ("Register " ++ pretty reg ++ " unmapped on indirect read")%string
1468+ $ TState.read_sreg_indirect ts reg;
1469+ mchoosel valvs
1470+ else
1471+ othrow
1472+ ("Register " ++ pretty reg ++ " unmapped; cannot read")%string
1473+ $ TState.read_reg ts reg.
1474+
1475+ Definition run_reg_trans_read (reg : reg) (racc : reg_acc)
1476+ (trs : IIS.TransRes.t) :
1477+ Exec.t TState.t string (reg_type reg * view) :=
1478+ guard_or "Translation read during the translation should be implicit"
1479+ (¬ (is_Some racc));;
1480+ root ← othrow "Could not find the translation root: error in translation assumptions"
1481+ (trs.(IIS.TransRes.root));
1482+ if decide (projT1 root = reg) is left eq then
1483+ mret (ctrans eq (projT2 root), 0%nat)
1484+ else
1485+ ts ← mGet;
1486+ othrow
1487+ ("Register " ++ pretty reg ++ " unmapped; cannot read")%string
1488+ $ TState.read_reg ts reg.
1489+
14371490(** Run a RegRead outcome.
14381491 Returns the register value based on the type of register and the access type. *)
14391492Definition run_reg_read (reg : reg) (racc : reg_acc) :
14401493 Exec.t (TState.t * IIS.t) string (reg_type reg) :=
1441- ts ← mget fst;
14421494 '(val, view) ←
1443- (if decide (reg ∈ relaxed_regs) then
1444- if decide (is_Some racc)
1445- then othrow
1446- ("Register " ++ pretty reg ++ " unmapped on direct read")%string
1447- $ TState.read_sreg_direct ts reg
1448- else
1449- valvs ← othrow
1450- ("Register " ++ pretty reg ++ " unmapped on indirect read")%string
1451- $ TState.read_sreg_indirect ts reg;
1452- mchoosel valvs
1495+ (* Check if the register is read during the translation *)
1496+ iis ← mget snd;
1497+ if iis.(IIS.trs) is Some trs then
1498+ Exec.liftSt fst $ run_reg_trans_read reg racc trs
14531499 else
1454- othrow
1455- ("Register " ++ pretty reg ++ " unmapped; cannot read")%string
1456- $ TState.read_reg ts reg);
1500+ run_reg_general_read reg racc;
14571501 mset snd $ IIS.add view;;
14581502 mret val.
14591503
@@ -1729,26 +1773,27 @@ Definition new_invalidation_opt (ti_new ti_old : option nat) : bool :=
17291773 | None, Some _ => true
17301774 end .
17311775
1732- Definition is_new_entry (path : list val) (ti_new : option nat)
1733- (entries : list (list val * nat * option nat)) : bool :=
1776+ Definition is_new_entry (val_ttbr : val) ( path : list val) (ti_new : option nat)
1777+ (entries : list (val * list val * nat * option nat)) : bool :=
17341778 forallb
1735- (λ '(p, t, ti),
1736- negb(path =? p) || new_invalidation_opt ti_new ti) entries.
1779+ (λ '(v, p, t, ti),
1780+ negb (val_ttbr =? v) || negb (path =? p)
1781+ || new_invalidation_opt ti_new ti) entries.
17371782
17381783(* Snapshots are sorted in the descending order of `trans_time`.
17391784 Thus, we do not have to use `trans_time` to check the interval *)
17401785Definition get_valid_entries_from_snapshots (snapshots : list (TLB.t * nat))
17411786 (mem : Memory.t)
17421787 (tid : nat)
17431788 (va : bv 64) (asid : bv 16) :
1744- result string (list (list val * nat * option nat)) :=
1789+ result string (list (val * list val * nat * option nat)) :=
17451790 fold_right (λ '(tlb, trans_time) entries,
17461791 candidates ← TLB.lookup mem tid tlb trans_time va asid;
17471792 prev ←@{result string} entries;
17481793 let new :=
1749- omap (λ '(path, ti_opt),
1750- if decide (is_new_entry path ti_opt prev) then
1751- Some (path, trans_time, ti_opt)
1794+ omap (λ '(val_ttbr, path, ti_opt),
1795+ if decide (is_new_entry val_ttbr path ti_opt prev) then
1796+ Some (val_ttbr, path, trans_time, ti_opt)
17521797 else None
17531798 ) candidates in
17541799 mret (new ++ prev)
@@ -1760,7 +1805,7 @@ Definition get_invalid_entries_from_snapshots (snapshots : list (TLB.t * nat))
17601805 (mem : Memory.t)
17611806 (tid : nat) (is_ets2 : bool)
17621807 (va : bv 64) (asid : bv 16) (ttbr : reg) :
1763- result string (list (list val * nat * option nat)) :=
1808+ result string (list (val * list val * nat * option nat)) :=
17641809 fold_right (λ '(tlb, trans_time) entries,
17651810 if decide (is_ets2 ∧ trans_time < ts.(TState.vwr) ⊔ ts.(TState.vrd)) then
17661811 entries
@@ -1770,9 +1815,9 @@ Definition get_invalid_entries_from_snapshots (snapshots : list (TLB.t * nat))
17701815 ts init mem tid tlb trans_time va asid ttbr;
17711816 prev ←@{result string} entries;
17721817 let new :=
1773- omap (λ '(path, ti_opt),
1774- if decide (is_new_entry path ti_opt prev) then
1775- Some (path, trans_time, ti_opt)
1818+ omap (λ '(val_ttbr, path, ti_opt),
1819+ if decide (is_new_entry val_ttbr path ti_opt prev) then
1820+ Some (val_ttbr, path, trans_time, ti_opt)
17761821 else None
17771822 ) candidates in
17781823 mret (new ++ prev)
@@ -1794,6 +1839,7 @@ Definition run_trans_start (trans_start : TranslationStartInfo)
17941839 (* lookup (successful results or faults) *)
17951840 let asid := trans_start.(TranslationStartInfo_asid) in
17961841 let va : bv 64 := trans_start.(TranslationStartInfo_va) in
1842+
17971843 trans_res ←
17981844 if decide (va_in_range va) then
17991845 ttbr ← mlift $ ttbr_of_regime va trans_start.(TranslationStartInfo_regime);
@@ -1804,15 +1850,25 @@ Definition run_trans_start (trans_start : TranslationStartInfo)
18041850 invalid_entries ← mlift $
18051851 get_invalid_entries_from_snapshots snapshots ts init mem tid is_ets2 va asid ttbr;
18061852 (* update IIS with either a valid translation result or a fault *)
1807- let valid_trs :=
1808- map (λ '(path, t, ti),
1809- IIS.TransRes.make (va_to_vpn va) t path ti) valid_entries in
1810- let invalid_trs :=
1811- map (λ '(path, t, ti),
1812- IIS.TransRes.make (va_to_vpn va) t path ti) invalid_entries in
1853+ valid_trs ←
1854+ for (val_ttbr, path, t, ti) in valid_entries do
1855+ val_ttbr ← othrow
1856+ "TTBR value type does not match with the value from the translation"
1857+ (val_to_regval ttbr val_ttbr);
1858+ let root := (Some (existT ttbr val_ttbr)) in
1859+ mret $ IIS.TransRes.make (va_to_vpn va) t root path ti
1860+ end ;
1861+ invalid_trs ←
1862+ for (val_ttbr, path, t, ti) in invalid_entries do
1863+ val_ttbr ← othrow
1864+ "TTBR value type does not match with the value from the translation"
1865+ (val_to_regval ttbr val_ttbr);
1866+ let root := (Some (existT ttbr val_ttbr)) in
1867+ mret $ IIS.TransRes.make (va_to_vpn va) t root path ti
1868+ end ;
18131869 mchoosel (valid_trs ++ invalid_trs)
18141870 else
1815- mret $ IIS.TransRes.make (va_to_vpn va) vpre_t [] None;
1871+ mret $ IIS.TransRes.make (va_to_vpn va) vpre_t None [] None;
18161872 mset PPState.iis $ IIS.set_trs trans_res.
18171873
18181874Definition read_fault_vpre (is_acq : bool)
0 commit comments