aboutsummaryrefslogtreecommitdiffstats
path: root/backend/Linear.v
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2019-09-05 11:19:22 +0200
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2019-09-05 11:19:22 +0200
commitc4cc75dc6abcb0eee6f3288e96fea4aec540fd68 (patch)
treee63d0437c3f3cb43eb26cced8ad94acc62e82dad /backend/Linear.v
parent7042070a3668ae149ec6a490b8e7c1a6aa82d6fe (diff)
downloadcompcert-kvx-c4cc75dc6abcb0eee6f3288e96fea4aec540fd68.tar.gz
compcert-kvx-c4cc75dc6abcb0eee6f3288e96fea4aec540fd68.zip
more proofs going through
Diffstat (limited to 'backend/Linear.v')
-rw-r--r--backend/Linear.v23
1 files changed, 20 insertions, 3 deletions
diff --git a/backend/Linear.v b/backend/Linear.v
index 447c6ba6..1443f795 100644
--- a/backend/Linear.v
+++ b/backend/Linear.v
@@ -28,7 +28,7 @@ Inductive instruction: Type :=
| Lgetstack: slot -> Z -> typ -> mreg -> instruction
| Lsetstack: mreg -> slot -> Z -> typ -> instruction
| Lop: operation -> list mreg -> mreg -> instruction
- | Lload: memory_chunk -> addressing -> list mreg -> mreg -> instruction
+ | Lload: trapping_mode -> memory_chunk -> addressing -> list mreg -> mreg -> instruction
| Lstore: memory_chunk -> addressing -> list mreg -> mreg -> instruction
| Lcall: signature -> mreg + ident -> instruction
| Ltailcall: signature -> mreg + ident -> instruction
@@ -160,11 +160,28 @@ Inductive step: state -> trace -> state -> Prop :=
step (State s f sp (Lop op args res :: b) rs m)
E0 (State s f sp b rs' m)
| exec_Lload:
- forall s f sp chunk addr args dst b rs m a v rs',
+ forall s f sp trap chunk addr args dst b rs m a v rs',
eval_addressing ge sp addr (reglist rs args) = Some a ->
Mem.loadv chunk m a = Some v ->
rs' = Locmap.set (R dst) v (undef_regs (destroyed_by_load chunk addr) rs) ->
- step (State s f sp (Lload chunk addr args dst :: b) rs m)
+ step (State s f sp (Lload trap chunk addr args dst :: b) rs m)
+ E0 (State s f sp b rs' m)
+ | exec_Lload_notrap1:
+ forall s f sp chunk addr args dst b rs m rs',
+ eval_addressing ge sp addr (reglist rs args) = None ->
+ rs' = Locmap.set (R dst)
+ (default_notrap_load_value chunk)
+ (undef_regs (destroyed_by_load chunk addr) rs) ->
+ step (State s f sp (Lload NOTRAP chunk addr args dst :: b) rs m)
+ E0 (State s f sp b rs' m)
+ | exec_Lload_notrap2:
+ forall s f sp chunk addr args dst b rs m a rs',
+ eval_addressing ge sp addr (reglist rs args) = Some a ->
+ Mem.loadv chunk m a = None ->
+ rs' = Locmap.set (R dst)
+ (default_notrap_load_value chunk)
+ (undef_regs (destroyed_by_load chunk addr) rs) ->
+ step (State s f sp (Lload NOTRAP chunk addr args dst :: b) rs m)
E0 (State s f sp b rs' m)
| exec_Lstore:
forall s f sp chunk addr args src b rs m m' a rs',